diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy deleted file mode 100644 index 23bb7049c30..00000000000 --- a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy +++ /dev/null @@ -1,4061 +0,0 @@ -// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" -// RUN: %diff "%s.expect" "%t" - -module Misc { - //Should not verify, as ghost loops should not be allowed to diverge. - method GhostDivergentLoop() - { - var a := new int [2]; - a[0] := 1; - a[1] := -1; - ghost var i := 0; - while (i < 2) - decreases * // error: not allowed on a ghost loop - invariant i <= 2 - invariant (forall j :: 0 <= j && j < i ==> a[j] > 0) - { - i := 0; - } - assert a[1] != a[1]; // ...for then this would incorrectly verify - } - - method ManyIndices(a: array3, b: array, m: int, n: int) - { - // the following invalid expressions were once incorrectly resolved: - var x := a[m, n]; // error - var y := a[m]; // error - var z := b[m, n, m, n]; // error - } - - method SB(b: array2, s: int) returns (x: int, y: int) - requires b != null - { - while - { - case b[x,y] == s => - } - } - - // -------- name resolution - - class Global { - var X: int - function F(x: int): int { x } - static ghost function G(x: int): int { x } - method M(x: int) returns (r: int) - { - r := x + X; - } - static method N(x: int) returns (r: int) - { - r := x + X; // error: cannot access instance field X from static method - } - } - - method TestNameResolution0() { - var z: int; - z := Global.X; // error: X is an instance field - z := F(2); // error: cannot resolve F - z := Global.F(2); // error: invocation of instance function requires an instance - z := G(2); // error: cannot resolve G - z := Global.G(2); - z := M(2); // error: cannot resolve M - z := Global.M(2); // error: call to instance method requires an instance - z := N(1); // error: cannot resolve N - z := Global.N(1); - - z := z(5); // error: using local as if it were a function - z := Global.z; // error: class Global does not have a member z - - var Global: Global; // a local variable with the name 'Global' - z := Global.X; // this means the instance field X of the object stored in the local variable 'Global' - var gg: Global := null; - var y := gg.G(5); - y := gg.N(5); - } - - datatype Abc = Abel | Benny | Cecilia(y: int) | David(x: int) | Eleanor - datatype Xyz = Alberta | Benny | Constantine(y: int) | David(x: int) - datatype Rst = David(x: int, y: int) - - ghost function Tuv(arg0: Abc, arg1: bool): int { 10 } - - class EE { - var Eleanor: bool - method TestNameResolution1() { - var a0 := Abel; - var a1 := Alberta; - var b0 := Benny; // error: there's more than one constructor with the name Benny; needs qualification - var b1 := Abc.Benny; - var b2 := Xyz.Benny; - var Benny := 15; // introduce a local variable with the name 'Benny' - var b3 := Benny; - var d0 := David(20); // error: constructor name David is ambiguous - var d1 := David; // error: constructor name David is ambiguous (never mind that the signature does - // not match either of them) - var d2 := David(20, 40); // error: constructor name Davis is ambiguous (never mind that the given - // parameters match the signature of only one of those constructors) - var d3 := Abc.David(20, 40); // error: wrong number of parameters - var d4 := Rst.David(20, 40); - var e := Eleanor; // this resolves to the field, not the Abc datatype constructor - assert Tuv(Abc.Eleanor, e) == 10; - } - } -} - -// --------------- ghost tests ------------------------------------- - -module HereAreMoreGhostTests { - datatype GhostDt = - Nil(ghost extraInfo: int) | - Cons(data: int, tail: GhostDt, ghost moreInfo: int) - - class GhostTests { - method M(dt: GhostDt) returns (r: int) { - ghost var g := 5; - r := g; // error: RHS is ghost, LHS is not - r := F(18, g); // error: RHS is a ghost and will not be available at run time - r := G(20, g); // it's fine to pass a ghost as a parameter to a non-ghost, because - // only the ghost goes away during compilation - r := N(22, g); // ditto - r := N(g, 22); // error: passing in 'g' as non-ghost parameter - r := P(24, 22); // error: 'P' is ghost, but its result is assigned to a non-ghost - - match (dt) { - case Nil(gg) => - case Cons(dd, tt, gg) => - r := G(dd, dd); // fine - r := G(dd, gg); // fine - r := G(gg, gg); // error: cannot pass ghost 'gg' as non-ghost parameter to 'G' - } - var dd; - dd := GhostDt.Nil(g); // fine - dd := GhostDt.Cons(g, dt, 2); // error: cannot pass 'g' as non-ghost parameter - ghost var dtg := GhostDt.Cons(g, dt, 2); // fine, since result is ghost - } - ghost function F(x: int, y: int): int { - y - } - function G(x: int, ghost y: int): int { - y // error: cannot return a ghost from a non-ghost function - } - function H(dt: GhostDt): int { - match dt - case Nil(gg) => gg // error: cannot return a ghost from a non-ghost function - case Cons(dd, tt, gg) => dd + gg // error: ditto - } - method N(x: int, ghost y: int) returns (r: int) { - r := x; - } - ghost method P(x: int, y: int) returns (r: int) { - ghost var g := 5; - r := y; // allowed, since the entire method is ghost - r := r + g; // fine, for the same reason - r := N(20, 20); // error: call to non-ghost method from ghost method is not okay - } - ghost method BreaksAreFineHere(t: int) - { - var n := 0; - ghost var k := 0; - while (true) - invariant n <= 112 - decreases 112 - n - { - label MyStructure: { - if (k % 17 == 0) { break MyStructure; } // this is fine, because it's a ghost method - k := k + 1; - } - label MyOtherStructure: - if (k % 17 == 0) { - break MyOtherStructure; - } else { - k := k + 1; - } - - if (n == 112) { - break; - } else if (n == t) { - return; - } - n := n + 1; - } - } - method BreakMayNotBeFineHere(ghost t: int) - { - var n := 0; - ghost var k := 0; - var p := 0; - while (true) - invariant n <= 112 - decreases 112 - n - { - label MyStructure: { - k := k + 1; - } - label MyOtherStructure: - if (k % 17 == 0) { - break MyOtherStructure; // this break is fine - } else { - k := k + 1; - } - - var dontKnow; - if (n == 112) { - ghost var m := 0; - label LoopLabel0: - label LoopLabel1: - while (m < 200) { - if (m % 103 == 0) { - if { - case true => break; // fine, since this breaks out of the enclosing ghost loop - case true => break LoopLabel0; // fine - case true => break LoopLabel1; // fine - } - } else if (m % 101 == 0) { - } - m := m + 3; - } - break; - } else if (dontKnow == 708) { - var q := 0; - while (q < 1) { - label IfNest: - if (p == 67) { - break break; // fine, since this is not a ghost context - } - q := q + 1; - } - } else if (n == t) { - } - n := n + 1; - p := p + 1; - } - } - method BreakMayNotBeFineHere_Ghost(ghost t: int) - { - var n := 0; - ghost var k := 0; - var p := 0; - while (true) - invariant n <= 112 - decreases 112 - n - { - label MyStructure: { - if (k % 17 == 0) { break MyStructure; } // error: break from ghost to non-ghost point - k := k + 1; - } - label MyOtherStructure: - if (k % 17 == 0) { - break MyOtherStructure; // this break is fine - } else { - k := k + 1; - } - - var dontKnow; - if (n == 112) { - ghost var m := 0; - label LoopLabel0: - label LoopLabel1: - while (m < 200) { - if (m % 103 == 0) { - if { - case true => break; // fine, since this breaks out of the enclosing ghost loop - case true => break LoopLabel0; // fine - case true => break LoopLabel1; // fine - } - } else if (m % 101 == 0) { - break break; // error: break out of non-ghost loop from ghost context - } - m := m + 3; - } - break; - } else if (dontKnow == 708) { - var q := 0; - while (q < 1) { - label IfNest: - if (p == 67) { - break break; // fine, since this is not a ghost context - } else if (*) { - break break; // fine, since this is not a ghost context - } else if (k == 67) { - break break; // error, because this is a ghost context - } - q := q + 1; - } - } else if (n == t) { - return; // error: this is a ghost context trying to return from a non-ghost method - } - n := n + 1; - p := p + 1; - } - } - } -} // HereAreMoreGhostTests - -module MiscMore { - method DuplicateLabels(n: int) { - var x; - if (n < 7) { - label DuplicateLabel: x := x + 1; - } else { - label DuplicateLabel: x := x + 1; - } - label DuplicateLabel: x := x + 1; - label DuplicateLabel': { - label AnotherLabel: - label DuplicateLabel': // error: duplicate label (shadowed by enclosing label) - label OneMoreTime: - x := x + 1; - } - label DuplicateLabel'': - label DuplicateLabel'': // error: duplicate label (shadowed by enclosing label) - x := x + 1; - label DuplicateLabel'': x := x + 1; // error: duplicate label (shadowed by dominating label) - } - - // --------------- constructors ------------------------------------- - // Note, more tests in module ClassConstructorTests below - - class ClassWithConstructor { - var y: int - method NotTheOne() { } - constructor InitA() { } - constructor InitB() modifies this { y := 20; } // error: don't use "this" in modifies of constructor - } - - method ConstructorTests() - { - var c := new ClassWithConstructor.InitB(); - c.InitB(); // error: not allowed to call constructors except during allocation - } - - // ------------------- datatype destructors --------------------------------------- - - datatype DTD_List = DTD_Nil | DTD_Cons(Car: int, Cdr: DTD_List, ghost g: int) - - method DatatypeDestructors(d: DTD_List) { - if { - case d.DTD_Nil? => - assert d == DTD_Nil; - case d.DTD_Cons? => - var hd := d.Car; - var tl := d.Cdr; - assert hd == d.Cdr; // type error - assert tl == d.Car; // type error - assert d.DTD_Cons? == d.Car; // type error - assert d == DTD_Cons(hd, tl, 5); - ghost var g0 := d.g; // fine - } - } -} // MiscMore - -// ------------------- print statements --------------------------------------- - -module GhostPrintAttempts { - method PrintOnlyNonGhosts(a: int, ghost b: int) - { - print "a: ", a, "\n"; - print "b: ", b, "\n"; // error: print statement cannot take ghosts - } -} - -// ------------------- auto-added type arguments ------------------------------ - -module MiscEvenMore { - class GenericClass { var data: T } - - method MG0(a: GenericClass, b: GenericClass) - requires a != null && b != null - modifies a - { - a.data := b.data; // allowed, since both a and b get the same auto type argument - } - - method G_Caller() - { - var x := new GenericClass; - MG0(x, x); // fine - var y := new GenericClass; - MG0(x, y); // also fine (and now y's type argument is constrained to be that of x's) - var z := new GenericClass; - var w := new GenericClass; - MG0(x, w); // this has the effect of making x's and y's type GenericClass - - y.data := z.data; // error: bool vs int - assert x.data == 5; // error: bool vs int - } - - datatype GList<+T> = GNil | GCons(hd: T, tl: GList) - - method MG1(l: GList, n: nat) - { - if (n != 0) { - MG1(l, n-1); - MG1(GCons(12, GCons(20, GNil)), n-1); - } - var t := GCons(100, GNil); // error: types don't match up (List<_T0> versus List) - t := GCons(120, l); // error: types don't match up (List<_T0> versus List) - } - - // ------------------- calc statements ------------------------------ - - method TestCalc(m: int, n: int, a: bool, b: bool) - { - calc { - a + b; // error: invalid line - n + m; - } - calc { - a && b; - n + m; // error: all lines must have the same type - } - calc ==> { - n + m; // error: ==> operator requires boolean lines - n + m + 1; - n + m + 2; - } - calc { - n + m; - n + m + 1; - ==> n + m + 2; // error: ==> operator requires boolean lines - } - } -} // MiscEvenMore - -module MyOwnModule { - class SideEffectChecks { - ghost var ycalc: int - - ghost method Mod(a: int) - modifies this - ensures ycalc == a - { - ycalc := a; - } - - ghost method Bad() - modifies this - ensures 0 == 1 - { - var x: int; - calc { - 0; - { Mod(0); } // error: methods with side-effects are not allowed - ycalc; - { ycalc := 1; } // error: heap updates are not allowed - 1; - { x := 1; } // error: updates to locals defined outside of the hint are not allowed - x; - { - var x: int; - x := 1; // this is OK - } - 1; - } - } - } -} - -// ------------------- nameless constructors ------------------------------ - -module MiscAgain { - class Y { - var data: int - constructor (x: int) - { - data := x; - } - constructor (y: bool) // error: duplicate constructor name - { - } - method Test() { - var i := new Y(5); - i := new Y(7); - - var s := new Luci.Init(5); - s := new Luci.FromArray(null); - s := new Luci(false); - s := new Luci(true); - - var l := new Lamb; - l := new Lamb(); // error: there is no default constructor - } - } - - class Luci { - constructor Init(y: int) { } - constructor (nameless: bool) { } - constructor FromArray(a: array) { } - method M() { } - } - - class Lamb { - method Jess() { } - method Gwen() { } - } - - // ------------------- assign-such-that and ghosts ------------------------------ - - method AssignSuchThatFromGhost() - { - var x: int; - ghost var g: int; - - x := *; - assume x == g; // this mix of ghosts and non-ghosts is cool (but, of course, - // the compiler will complain) - - x :| assume x == g; // this is cool, since it's an assume (but, of course, the - // compiler will complain) - - x :| x == 5; - g :| g <= g; - g :| assume g < g; // the compiler will complain here, despite the LHS being - // ghost -- and rightly so, since an assume is used - } -} // MiscAgain - -// ------------------------ inferred type arguments ---------------------------- - -// Put the following tests in a separate module, so that the method bodies will -// be type checked even if there are resolution errors in other modules. -module NoTypeArgs0 { - datatype List<+T> = Nil | Cons(T, List) - datatype Tree<+A,+B> = Leaf(A, B) | Node(Tree, Tree) - - method DoAPrefix0(xs: List) returns (ys: List) - { - ys := xs; - } - - method DoAPrefix1(xs: List) returns (ys: List) - { - ys := xs; // error: List cannot be assign to a List - } - - method DoAPrefix2(xs: List) returns (ys: List) - { - ys := xs; // error: List cannot be assign to a List - } - - ghost function FTree0(t: Tree): Tree - { - match t - case Leaf(_,_) => t - case Node(x, y) => x - } - - ghost function FTree1(t: Tree): Tree - { - match t - case Leaf(_,_) => t - case Node(x, y) => y // error: y does not have the right type - } - - ghost function FTree2(t: Tree): Tree - { - t - } -} - -module NoTypeArgs1 { - datatype Tree = Leaf(A, B) | Node(Tree, Tree) - - ghost function FTree3(t: Tree): Tree // error: type of 't' does not have enough type parameters - { - t - } -} - -// ----------- let-such-that expressions ------------------------ - -module MiscMisc { - method LetSuchThat(ghost z: int, n: nat) - { - var x: int; - x := var y :| y < 0; y; // fine for the resolver (but would give a verification error for not being deterministic) - - x := var w :| w == 2*w; w; // fine (even for the verifier, this one) - x := var w := 2*w; w; // error: the 'w' in the RHS of the assignment is not in scope - ghost var xg := var w :| w == 2*w; w; - } -} - -// ------------ quantified variables whose types are not inferred ---------- - -module NonInferredType { - ghost predicate P(x: T) - - method InferredType(x: int) - { - var t; - assume forall z :: P(z) && z == t; - assume t == x; // this statement determines the type of t and z - } - - method NonInferredType(x: int) - { - var t; // error: the type of t is not determined - assume forall z :: P(z) && z == t; // error: the type of z is not determined - } -} - -// ------------ Here are some tests that lemma contexts don't allocate objects ------------- - -module GhostAllocationTests { - class G { } - iterator GIter() { } - class H { constructor () } - lemma GhostNew0() - ensures exists o: G :: fresh(o) - { - var p := new G; // error: lemma context is not allowed to allocate state - p := new G; // error: ditto - } - - method GhostNew1(n: nat, ghost g: int) returns (t: G, z: int) - { - if n < 0 { - z, t := 5, new G; // fine - } - if n < g { - var tt := new H(); // error: 'new' not allowed in ghost contexts - } - } - - method GhostNew2(ghost b: bool) - { - if (b) { - var y := new GIter(); // error: 'new' not allowed in ghost contexts (and a non-ghost method is not allowed to be called here either) - } - } -} -module MoreGhostAllocationTests { - class G { } - method GhostNew3(n: nat) { - var g := new G; - calc { - 5; - { var y := new G; } // error: 'new' not allowed in lemma contexts - 2 + 3; - } - } - ghost method GhostNew4(g: G) - modifies g - { - } -} - -module NewForallAssign { - class G { } - method NewForallTest(n: nat) { - var a := new G[n]; - forall i | 0 <= i < n { - a[i] := new G; // error: 'new' is currently not supported in forall statements - } } -} -module NewForallProof { - class G { } - method NewForallTest(n: nat) { var a := new G[n]; - forall i | 0 <= i < n ensures true { // this makes the whole 'forall' statement into a ghost statement - a[i] := new G; // error: proof-forall cannot update state (and 'new' not allowed in ghost contexts, but that's checked at a later stage) - } } -} - -// ------------------------- underspecified types ------------------------------ - -module UnderspecifiedTypes { - method M(S: set) { - var n, p, T0 :| 12 <= n && n in T0 && 10 <= p && p in T0 && T0 <= S && p % 2 != n % 2; - var T1 :| 12 in T1 && T1 <= S; - var T2 :| T2 <= S && 12 in T2; - var T3 :| 120 in T3; // error: underspecified type - var T3'0: set :| 120 in T3'0; - var T3'1: multiset :| 120 in T3'1; - var T3'2: map :| 120 in T3'2; - var T3'3: seq :| 120 in T3'3; - var T3'4: bool :| 120 in T3'4; // error: second argument to 'in' cannot be bool - var T4 :| T4 <= S; - } -} - -// ------------------------- lemmas ------------------------------ -module MiscLemma { - class L { } - - // a lemma is allowed to have out-parameters, but not a modifies clause - lemma MyLemma(x: int, l: L) returns (y: int) - requires 0 <= x - modifies l - ensures 0 <= y - { - y := x; - } -} - -// ------------------------- statements in expressions ------------------------------ - -module StatementsInExpressions { - class MyClass { - ghost method SideEffect() - modifies this - { - } - - method NonGhostMethod() - { - } - - ghost function F(): int - { - calc { - 6; - { assert 6 < 8; } - { var x := 8; - while x != 0 - decreases * // error: cannot use 'decreases *' here - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - 6; - } - 5 - } - - var MyField: int - ghost var MyGhostField: int - - method N() - { - var y := - calc { - 6; - { assert 6 < 8; } - { var x := 8; - while x != 0 - decreases * // error: cannot use 'decreases *' here - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - 6; - } - 5; - } - } -} - -module StmtExprOutParams { - - lemma MyLemma() - - lemma OutParamLemma() returns (y: int) - - ghost function UseLemma(): int - { - MyLemma(); - OutParamLemma(); // error: has out-parameters - 10 - } -} - -module GhostLetExpr { - method M() { - ghost var y; - var x; - var g := G(x, y); - ghost var h := ghost var ta := F(); 5; - var j; j := var tb := F(); 5; // fine (tb is ghost, j is not, but RHS body doesn't depend on tb) - assert h == j; - } - - ghost function F(): int - { 5 } - - function G(x: int, ghost y: int): int - { - assert y == x; - y // error: not allowed in non-ghost context - } - - datatype Dt = MyRecord(a: int, ghost b: int) - - method P(dt: Dt) { - match dt { - case MyRecord(aa, bb) => - ghost var z := aa + F(); - ghost var t0 := var y := z; z + 3; - ghost var t1 := ghost var y := z + bb; y + z + 3; - var t2; t2 := ghost var y := z; y + 3; // t2 is not ghost - error - } - } - - function FM(e: bool): int - { - if e then - G(5, F()) - else - var xyz := F(); // fine, because 'xyz' becomes ghost automatically - G(5, xyz) - } -} - -module ObjectType { - type B - datatype Dt = Blue | Green - codatatype CoDt = Cons(int, CoDt) - class MyClass { } - - method M(zz: array, j: int, b: B, co: CoDt, g: G) returns (o: object) - requires zz != null && 0 <= j < zz.Length - { - o := b; // error - o := 17; // error - o := zz[j]; // error - o := null; - o := zz; - o := new MyClass; - o := o; - o := g; // error - o := Blue; // error - o := co; // error - } -} -// ------------------ modify statment --------------------------- -module MiscModify { - class ModifyStatementClass { - var x: int - ghost var g: int - method M() { - modify x; // error: type error - } -} } -module MiscModifiesGhost { - class ModifyStatementClass { - var x: int - ghost var g: int - ghost method G0() - modifies `g - modifies `x // error: non-ghost field mentioned in ghost context -} } -module ModifyStatementClass_More { - class C { - var x: int - ghost var g: int - ghost method G0() - modifies `g - { - modify `g; - modify `x; // error: non-ghost field mentioned in ghost context - } - method G1() - modifies this - { - modify `x; - if g < 100 { - // we are now in a ghost context - modify `x; // error: non-ghost field mentioned in ghost context - } - } - method G2(y: nat) - modifies this - { - if g < 100 { - // we're now in a ghost context - var n := 0; - while n < y - modifies `x // error: non-ghost field mentioned in ghost context - { - if * { - g := g + 1; // if we got as far as verification, this would be flagged as an error too - } - n := n + 1; - } - } - modify `x; // fine - ghost var i := 0; - while i < y - modifies `x // error: non-ghost field mentioned in ghost context - { - i := i + 1; - } - } - } -} - -module LhsLvalue { - method M() - { - var mySeq: seq; - var a := new int[78]; - var b := new int[100, 200]; - var c := new MyRecord[29]; - - mySeq[0] := 5; // error: cannot assign to a sequence element - mySeq[0] := MyMethod(); // error: ditto - a[0] := 5; - a[0] := MyMethod(); - b[20, 18] := 5; - b[20, 18] := MyMethod(); - c[25].x := 5; // error: cannot assign to a destructor - c[25].x := MyMethod(); // error: ditto - mySeq[0..4] := 5; // error: cannot assign to a range - mySeq[0..4] := MyMethod(); // error: ditto - a[0..4] := 5; // error: cannot assign to a range - a[0..4] := MyMethod(); // error: ditto - } - - datatype MyRecord = Make(x: int, y: int) - - method MyMethod() returns (w: int) -} - -// ------------------- dirty loops ------------------- -module MiscEtc { - method DirtyM(S: set) { - forall s | s in S ensures s < 0 - assert s < 0; // error: s is unresolved - } - - // ------------------- tuples ------------------- - - method TupleResolution(x: int, y: int, r: real) - { - var unit: () := (); - var expr: int := (x); - var pair: (int,int) := (x, x); - var triple: (int,int,int) := (y, x, x); - var badTriple: (int,real,int) := (y, x, r); // error: parameters 1 and 2 have the wrong types - var quadruple: (int,real,int,real) := (y, r, x); // error: trying to use a triple as a quadruple - - assert unit == (); - assert pair.0 == pair.1; - assert triple.2 == x; - - assert triple.2; // error: 2 has type int, not the expected bool - assert triple.3 == pair.x; // error(s): 3 and x are not destructors - - var k0 := (5, (true, 2, 3.14)); - var k1 := (((false, 10, 2.7)), 100, 120); - if k0.1 == k1.0 { - assert false; - } else if k0.1.1 < k1.0.1 { - assert k1.2 == 120; - } - - // int and (int) are the same type (i.e., there are no 1-tuples) - var pp: (int) := x; - var qq: int := pp; - } - - // ------------------- conversions ------------------- - - method TypeConversions(m: nat, i: int, r: real) returns (n: nat, j: int, s: real) - { - n := r as int; - j := r as int; - s := m as real; // nat->real is allowed, just like int->real is - s := i as real; - s := i as real / 2; // error: division expects two reals - s := 15 % s; // error: modulus is not defined for reals - - s := (2.0 / 1.7) + (r / s) - (--r) * -12.3; - - s := s as real; // fine (identity transform) - j := j as int; // fine (identity transform) - j := n as int; // fine (identity transform) - } -} - -// --- filling in type arguments and checking that there aren't too many --- - -module TypeArgumentCount { - class C { - var f: T - } - - method R0(a: array3, c: C) - - method R1() - { - var a: array3; - var c: C; - } - - method R2() - { - var a: array3; // error: too many type arguments - var c: C; // error: too many type arguments - } -} - -// --- Type synonyms --- - -module BadTypeSynonyms { - datatype List = Nil | Cons(T, List) - type BadSyn0 = List // error: must have at least one type parameter - type BadSyn1 = badName // error: badName does not denote a type - type BadSyn2 = List // error: X does not denote a type - type BadSyn2 = int // error: repeated name -} - -// --- cycles --- - -module CycleError0 { - type A = A // error: cycle: A -> A -} -module CycleError1 { - type A = B // error: cycle: A -> B -> A - type B = A -} -module CycleError2 { - type A = B // error: cycle: A -> B -> A - type B = set -} -module CycleErrors3 { - type A = (B, D) - type B = C - class C { // error: since A is not auto-init, class C needs a constructor - var a: A // this is fine - } - datatype D = Make(A, B, C) // warning: D is empty -} -module CycleError4 { - type A = B // error: cycle: A -> B -> A - type B = C - class C { } -} -module CycleError5 { - type A = B // error: cycle: A -> B -> A - type B = Dt - datatype Dt = Make(T) -} -module CycleError6 { - type A = Dt> // error: cycle A -> A - datatype Dt = Make(T) -} - -// --- attributes in top-level declarations --- - -module MiscIterator { - iterator {:myAttribute x} Iter() { // error: x does not refer to anything - } - - class {:myAttribute x} C { // error: x does not refer to anything - } - - datatype {:myAttribute x} Dt = Blue // error: x does not refer to anything - - type {:myAttribute x} Something // error: x does not refer to anything - - type {:myAttribute x} Synonym = int // error: x does not refer to anything -} - -module {:myAttribute x} Modulette { // error: x does not refer to anything -} - -// --- abstract types with type parameters --- - -module OpaqueTypes0 { - type P - method M(p: P) returns (q: P) // error: wrong param count - { - q := p; - } -} - -module OpaqueTypes1 { - type P - - method M0(p: P) returns (q: P) - { - q := p; - var m: P; // error: BX undefined - } - - method M1(p: P) returns (q: P) // type parameter of q's type inferred - { - q := p; - } - - method M2(p: P) returns (q: P) - { - q := p; // error: cannot assign P to P - } - - method M3(p: P) returns (q: P) - { - q := p; // error: cannot assign P to P - } - - method M4() returns (p: P, q: P) - { - q := p; // error: cannot assign P to P - p := q; // error: cannot assign P to P - } - - method EqualityTests(p: P, q: P, r: P) - { - assert p != r; // error: types must be the same in order to do compare - assert q != r; // error: types must be the same in order to do compare - assert p != q; // error: types must be the same in order to do compare - } -} - -// ----- new trait ------------------------------------------- - -module MiscTrait { - trait J { } - type JJ = J - method TraitSynonym() - { - var x := new JJ; // error: new cannot be applied to a trait - } -} - -// ----- set comprehensions where the term type is finite ----- - -module ObjectSetComprehensionsNever { - // the following set comprehensions are known to be finite - ghost function A() : set { set o : object | true :: o } // error: a function is not allowed to depend on the allocated state - function B() : set { set o : object | true :: o } // error: a function is not allowed to depend on the allocated state -} -module ObjectSetComprehensionsSometimes { - // outside functions, the comprehension is permitted, but it cannot be compiled - lemma C() { var x; x := set o : object | true :: o; } - - method D() { var x; x := set o : object | true :: o; } // error: not (easily) compilable, so this is allowed only in ghost contexts -} - -// ------ regression test for type checking of integer division ----- - -module MiscTests { - method IntegerDivision(s: set) - { - var t := s / s; // error: / cannot be used with sets - } - - // ----- decreases * tests ---- - - method NonTermination_A() - { - NonTermination_B(); // error: to call a non-terminating method, the caller must be marked 'decreases *' - } - - method NonTermination_B() - decreases * - { - while true - decreases * - { - } - } - - method NonTermination_C() - { - while true - decreases * // error: to use an infinite loop, the enclosing method must be marked 'decreases *' - { - } - } - - method NonTermination_D() - decreases * - { - var n := 0; - while n < 100 // note, no 'decreases *' here, even if the nested loop may fail to terminate - { - while * - decreases * - { - } - n := n + 1; - } - } -} - -// ------------ type variables whose values are not inferred ---------- - -module NonInferredTypeVariables { - class C { - var f: CT - } - - predicate P(x: int) - { - x < 100 - } - ghost function Q(x: int): QT - { - var qt :| true; qt - } - method M(n: nat) - { - var a := new MT[n]; - } - method N(n: nat) returns (x: NT) - { - var a := new NT[10]; - x := a[3]; - } - - method DeterminedClient(n: nat) - { - ghost var q := Q(n); - var x := N(n); - var a := new array; - var c := new C; - var s: set; - var ss := new set[15]; - - q := 3.14; // this will determine the type parameter of Q to be 'real' - x := 3.14; // this will determine the type parameter of N to be 'real' - if a.Length != 0 { - a[0] := 3.14; // this will determine the type parameter of 'array' to be 'real' - } - c.f := 3.14; // this will determine the type parameter of 'C' to be 'real' - var containsPi := 3.14 in s; // this will determine the type parameter of 'set' to be 'real' - ss[12] := s; // this will determine the type parameter of 'array>' to be 'real' - } - method BadClient(n: nat) - { - var p := P(n); // error: cannot infer the type argument for P - ghost var q := Q(n); // error: cannot infer the type argument for Q (and thus q's type cannot be determined either) - M(n); // error: cannot infer the type argument for M - var x := N(n); // error: cannot infer the type argument for N (and thus x's type cannot be determined either) - var a := new array; // error: cannot infer the type argument for 'array' - var c := new C; // error: cannot infer the type argument for 'C' - var s: set; // type argument for 'set' - var ss := new set[15]; // error: cannot infer the type argument in 'array>' - var what; // error: the type of this local variable in underspecified - } - method MoreBadClient() - { - var b0 := forall s :: s <= {} ==> s == {}; // error: type of s underspecified - var b1 := forall s: set :: s <= {} ==> s == {}; // error: type of s underspecified - var b2 := forall c: C? :: c in {null} ==> c == null; // error: type parameter of c underspecified - - // In the following, the type of the bound variable is completely determined. - var S: set>; - ghost var d0 := forall s :: s == {7} ==> s != {}; - var d1 := forall s: set :: s in S ==> s == {}; - var ggcc0: C; - var ggcc1: C; // error: full type cannot be determined - ghost var d2 := forall c: C? :: c != null ==> c.f == 10; - ghost var d2' := forall c: C? :: c == ggcc0 && c != null ==> c.f == 10; - ghost var d2'' := forall c: C? :: c == ggcc1 && c != null ==> c.f == c.f; // error: here, type of c is not determined - - var d0' := forall s :: s == {7} ==> s != {}; - var d0'' := forall s :: s <= {7} ==> s == {}; - } -} - -// -------------- signature completion ------------------ - -module SignatureCompletion { - // datatype signatures do not allow auto-declared type parameters on the LHS - datatype Dt = Ctor(X -> Dt) // error: X is not a declared type - datatype Et = Ctor(X -> Et, Y) // error: X is not a declared type - - - method My0(s: set, x: A -> B) - method My1(x: A -> B, s: set) - method My2(s: set, x: A -> B) - method My3(x: A -> B, s: set) - - ghost function F0(s: set, x: A -> B): int - ghost function F1(x: A -> B, s: set): int - ghost function F2(s: set, x: A -> B): int - ghost function F3(x: A -> B, s: set): int -} - -// -------------- more fields as frame targets -------------------- - -module FrameTargetFields { - class C { - var x: int - var y: int - ghost var z: int - - method M() - modifies this - { - var n := 0; - ghost var save := y; - while n < x - modifies `x - { - n, x := n + 1, x - 1; - } - assert y == save; - } - - ghost method N() - modifies this - modifies `y // resolution error: cannot mention non-ghost here - modifies `z // cool - { - } - } -} - -module FrameTargetFields_More { - class C { - var x: int - var y: int - ghost var z: int - method P() - modifies this - { - ghost var h := x; - while 0 <= h - modifies `x // resolution error: cannot mention non-ghost here - modifies `z // cool - { - h, z := h - 1, 5 * z; - } - } - } -} - -// ------------------------------------------------------ - -module AmbiguousModuleReference { - module A { - module Inner { - ghost predicate Q() - } - } - module B { - module Inner { - ghost predicate Q() - } - } - module OpenClient { - import opened A - import opened B - lemma M() { - var a := A.Inner.Q(); // fine - var b := B.Inner.Q(); // fine - var p := Inner.Q(); // error: Inner is ambiguous (A.Inner or B.Inner) - } - } -} - -// -------------------------------------------------- - -module GhostLet { - method M() { - var x: int; - x := ghost var tmp := 5; tmp; // error: ghost -> non-ghost - x := ghost var tmp := 5; 10; // fine - x := ghost var a0, a1 :| a0 == 0 && a1 == 1; a0 + a1; // error: ghost -> non-ghost - x := ghost var a :| 0 <= a; 10; // fine - } -} - -// ------------------- tuple equality support ------------------- - -module TupleEqualitySupport { - datatype GoodRecord = GoodRecord(set<(int,int)>) - datatype BadRecord = BadRecord(set<(int, int->bool)>) // error: this tuple type does not support equality -} - -// ------------------- non-type variable names ------------------- - -module NonTypeVariableNames { - type X = int - - module Y { } - - method M(m: map) - { - assert X == X; // error (x2): type name used as variable - assert Y == Y; // error (x2): module name used as variable - assert X in m; // error (x2): type name used as variable - assert Y in m; // error (x2): module name used as variable - } - - method N(k: int) - { - assert k == X; // error (x2): type name used as variable - assert k == Y; // error (x2): module name used as variable - X := k; // error: type name used as variable - Y := k; // error: module name used as variable - } -} - -// ------------------- assign-such-that and let-such-that ------------------- - -module SuchThat { - method M() { - var x: int; - x :| 5 + 7; // error: constraint should be boolean - x :| x; // error: constraint should be boolean - var y :| 4; // error: constraint should be boolean - } - ghost function F(): int { - var w :| 6 + 8; // error: constraint should be boolean - w - } -} - -// ---------------------- NEW STUFF ---------------------------------------- - -module GhostTests { - class G { } - - method GhostNew3(n: nat) - { - var g := new G; - calc { - 5; - 2 + 3; - { if n != 0 { GhostNew3(n-1); } } // error: cannot call non-ghost method in a ghost context - 1 + 4; - { GhostNew4(g); } // error: cannot call method with nonempty modifies - -5 + 10; - } - } - - ghost method GhostNew4(g: G) - modifies g - { - } - - class MyClass { - ghost method SideEffect() - modifies this - { - } - - method NonGhostMethod() - { - } - - ghost method M() - modifies this - { - calc { - 5; - { SideEffect(); } // error: cannot call method with side effects - 5; - } - } - ghost function F(): int - { - calc { - 6; - { assert 6 < 8; } - { NonGhostMethod(); } // error: cannot call non-ghost method - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - { MyField := 12; } // error: cannot assign to a field, and especially not a non-ghost field - { MyGhostField := 12; } // error: cannot assign to any field - { SideEffect(); } // error: cannot call (ghost) method with a modifies clause - { var x := 8; - while x != 0 - modifies this // error: cannot use a modifies clause on a loop inside a hint - { - x := x - 1; - } - } - 6; - } - 5 - } - var MyField: int - ghost var MyGhostField: int - method N() - { - var y := - calc { - 6; - { assert 6 < 8; } - { NonGhostMethod(); } // error: cannot call non-ghost method - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - { MyField := 12; } // error: cannot assign to a field, and especially not a non-ghost field - { MyGhostField := 12; } // error: cannot assign to any field - { M(); } // error: cannot call (ghost) method with a modifies clause - { var x := 8; - while x != 0 - modifies this // error: cannot use a modifies clause on a loop inside a hint - { - x := x - 1; - } - } - { var x := 8; - while x != 0 - { - x := x - 1; - } - } - 6; - } - 5; - } - } -} -module CallsInStmtExpr { - class MyClass { - lemma MyLemma() - ghost method MyEffectlessGhostMethod() - ghost function UseLemma(): int - { - MyEffectlessGhostMethod(); // error: cannot call ghost methods (only lemmas) from this context - MyLemma(); - 10 - } - } -} - -module EvenMoreGhostTests { - ghost method NiceTry() - ensures false - { - while (true) - decreases * // error: not allowed here - { - } - } - method BreakMayNotBeFineHere() - { - var n := 0; - var p := 0; - while (true) - { - var dontKnow; - if (n == 112) { - } else if (dontKnow == 708) { - while * { - label IfNest: - if (p == 67) { - break break; // fine, since this is not a ghost context - } else if (*) { - break break break; // error: tries to break out of more loop levels than there are - } - } - } - } - } -} - -module BadGhostTransfer { - datatype DTD_List = DTD_Nil | DTD_Cons(Car: int, Cdr: DTD_List, ghost g: int) - - method DatatypeDestructors_Ghost(d: DTD_List) { - var g1; g1 := d.g; // error: cannot use ghost member in non-ghost code - } - method AssignSuchThatFromGhost() - { - var x: int; - ghost var g: int; - - x := g; // error: ghost cannot flow into non-ghost - - x := *; - assume x == g; // this mix of ghosts and non-ghosts is cool (but, of course, - // the compiler will complain) - - x :| x == g; // error: left-side has non-ghost, so RHS must be non-ghost as well - - x :| assume x == g; // this is cool, since it's an assume (but, of course, the - // compiler will complain) - - x :| x == 5; - g :| g <= g; - g :| assume g < g; // the compiler will complain here, despite the LHS being - // ghost -- and rightly so, since an assume is used - } -} - -module MoreGhostPrintAttempts { - method TestCalc_Ghost(m: int, n: int, a: bool, b: bool) - { - calc { - n + m; - { print n + m; } // error: non-ghost statements are not allowed in hints - m + n; - } - } -} - -module MoreLetSuchThatExpr { - method LetSuchThat_Ghost(ghost z: int, n: nat) - { - var x; x := var y :| y < z; y; // error: contraint depend on ghost (z) - } -} - -module UnderspecifiedTypedShouldBeResolvedOnlyOnce { - method CalcTest0(s: seq) { - calc { - 2; - var t :| true; 2; // error: type of 't' is underspecified - } - } -} - -module LoopResolutionTests { - class C { - var x: int - ghost var y: int - } - - - ghost method M(c: C) - modifies c - { - var n := 0; - while n < 100 - modifies c`y - modifies c`x // error: not allowed to mention non-ghost field in modifies clause of ghost loops - { - c.x := c.x + 1; // error: assignment to non-ghost field not allowed here - } - } - - - method MM(c: C) - modifies c - { - var n := 0; - while - invariant n <= 100 - modifies c // regression test - { - case n < 100 => n := n + 1; - } - } - - - method MMX(c: C, ghost g: int) - modifies c - { - var n := 0; - while - invariant n <= 100 - modifies c`y - modifies c`x // error: not allowed to mention non-ghost field in modifies clause of ghost loops - { - case n < 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop - case g < 56 && n != 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop - } - } - - - method MD0(c: C, ghost g: nat) - modifies c - decreases * - { - var n := 0; - while n + g < 100 - invariant n <= 100 - decreases * // error: disallowed on ghost loops - { - n := n + 1; // error: cannot assign to non-ghost in a ghost loop - } - } - - - method MD1(c: C, ghost g: nat) - modifies c - decreases * - { - var n := 0; - while - invariant n <= 100 - decreases * // error: disallowed on ghost loops - { - case n + g < 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop - } - } -} - -module UnderspecifiedTypesInAttributes { - function P(x: T): int - method M() { - var {:myattr var u :| true; 6} v: int; // error: type of u is underspecified - var j {:myattr var u :| true; 6} :| 0 <= j < 100; // error: type of u is underspecified - - var a := new int[100]; - forall lp {:myattr var u :| true; 6} | 0 <= lp < 100 { // error: type of u is underspecified - a[lp] := 0; - } - - modify {:myattr P(10)} {:myattr var u :| true; 6} a; // error: type of u is underspecified - - calc {:myattr P(10)} {:myattr var u :| true; 6} // error: type of u is underspecified - { - 5; - } - } -} - -// ------------------- infer array types for Indexable and MultiIndexable XConstraints ---------- -// ------------------- using weaker subtyping constraints ---------- - -module AdvancedIndexableInference { - datatype MyRecord = Make(x: int, y: int) - method M(d: array, e: seq) - requires d.Length == 100 == |e| - { - if * { - var c := d; - var xx := c[25].x; - } else if * { - var c := d; - var xx := c[25..50][10].x; - } else if * { - var c := e; - var xx := c[25].x; - } else { - var c := e; - var xx := c[25..50][10].x; - } - } -} - -// -------------------------- - -module TypeConversions { - trait J { } - class C extends J { } - method M() returns (x: int, n: nat, o: object, j: J, c: C) { - n := x as nat; // yes, this is allowed now - o := j; - j := o; // OK for type resolution, but must be proved - j := o as J; // yes, this is allowed now - j := c; - c := j; // OK for type resolution, but must be proved - c := j as C; // yes, this is allowed now - var oo := o as realint; // error: there's no such type as "realint" (this once used to crash Dafny) - } -} - -// --------------------- regression - -module Regression_NewType { - class C { } - newtype MyInt = x: int | {} == set c: C | c // this once crashed Dafny -} - -// --------------------- regression - -module PrefixGeneratorDuplicates { - greatest predicate P() - greatest predicate P() // error: duplicate name (this once crashed Dafny) - greatest lemma L() - greatest lemma L() // error: duplicate name (this once crashed Dafny) -} - -// ------------------- unary TLA+ style predicates ------------------- - -module TLAplusOperators { - ghost function BadA(y: int): int // error: body has wrong return type - { - && 5 + y // error: using operator "&&" requires the operand to be boolean - } - ghost function BadB(y: int): bool - { - && 5 + y // error: using operator "&&" requires the operand to be boolean - } - ghost function BadC(y: int): int // error: body has wrong return type - { - || 5 + y // error: using operator "||" requires the operand to be boolean - } - ghost function BadD(y: int): bool - { - || 5 + y // error: using operator "||" requires the operand to be boolean - } - ghost function BadE(y: int): int // error: body has wrong return type - { - && (|| 5 + y) // error: bad types - } -} - -// ------------------------- divided constructors ------------------- - -module DividedConstructors { - - class MyClass { - var a: nat - var b: nat - var c: nat - var n: MyClass - const t := 17 - static const g := 25 - - constructor Init(x: nat) - { - this.a := this.b; // this use of "this" in RHS is allowed - ((this)).b := 10; - n := new MyClass(); - n.a := 10; // error: not allowed use of "this" in this way - c := a + b; // error (x2): not allowed "this" in RHS - var th := this; // error: not allowed "this" in RHS - Helper(); // error: not allowed to call instance method - var mc := new MyClass(); - StaticHelper(mc); - this.StaticHelper(mc); // error: "this" is benign here; yet, it seems consistent to disallow it - StaticHelper(this); // error: cannot use "this" here - P(a); // error: cannot use "this" here - P(g); - P(this.g); // error: "this" is benign here; yet, it seems consistent to disallow it - modify this; // error: cannot use "this" here - modify this`c; // error: cannot use "this" here - modify `c; // error: cannot use (implicit) "this" here - new; - a := a + b; - Helper(); - } - - method Helper() - { - } - - static method StaticHelper(mc: MyClass) - { - } - - static method P(x: nat) - { - } - - constructor () - { - a, c := 0, 0; - new; - } - - } -} - -module ConstructorsThisUsage { - class C { - var x: int - constructor M() - requires this != null // error: cannot use "this" here - modifies this // error: cannot use "this" here (but we just issue a deprecation warning) - decreases this.x // error: cannot use "this" here - ensures this.x == 5 - { - x := 5; - } - } -} - -module ReturnBeforeNew { - class C { - var a: int - var b: int - constructor TriesToReturnBeforeNew(xyz: int) - { - a := 0; - if xyz < 100 { - return; // error: "return" is not allowed before "new;" - } - } - } -} - -// ---------------- required auto-initialization ----------------------- - -module ZI { - // the following are different syntactic ways of saying that the type - // must support auto-initialization - type ZA(0) - type ZB(==)(0) - type ZC(0)(==) - type ZD(==,0) - type ZE(0,==) - type Y - - method P(x: G) - method M0(a: ZA, b: ZB, c: ZC, d: ZD, e: ZE, f: F, g: G, y: Y) - { - P(a); - P(b); - P(c); - P(d); - P(e); - P(f); // error: type of argument is expected to support auto-initialization - P(g); - P(y); // error: type of argument is expected to support auto-initialization - } - - datatype List = Nil | Cons(T, List) - method M1(xs: List, ys: List) { - P(xs); // yay, type of argument does support auto-initialization - P(ys); // yay, type of argument does support auto-initialization - } - - class Cls { - var q: int - var rs: List> - } - method M2(c: Cls?) { - P(c); - } - - newtype byte = x: int | 0 <= x < 256 // supports auto-initialization - newtype MyInt = int // supports auto-initialization - newtype SixOrMore = x | 6 <= x ghost witness 6 - newtype AnotherSixOrMore = s: SixOrMore | true ghost witness 6 - newtype MySixOrMore = x: MyInt | 6 <= x ghost witness 6 - // The resolver uses the presence/absence of a "witness" clause to figure out if the type - // supports auto-initialization. This can be inaccurate. If the type does not have a - // "witness" clause, some type replacements may slip by the resolver, but will then be - // caught by the verifier when the witness test is performed (because the witness test - // uses a zero value in the absence of a "witness" clause). - // A "ghost witness" clause tells the resolver that the type does not support - // auto-initialization, but only ghost auto-initialzation. - - newtype UnclearA = x: int | true ghost witness 0 // actually supports auto-initialization, but has a "ghost witness" clause - newtype UnclearB = x | x == 6 && x < 4 // "witness" clause omitted; type does not actually support auto-initialization - - method M3(a: byte, b: MyInt, c: SixOrMore, d: AnotherSixOrMore, e: MySixOrMore, - ua: UnclearA, ub: UnclearB) { - P(a); - P(b); - P(c); // error: type of argument is expected to support auto-initialization - P(d); // error: type of argument is expected to support auto-initialization - P(e); // error: type of argument is expected to support auto-initialization - P(ua); // error: as far as the resolver can tell, type of argument does not support auto-initialization - P(ub); // fine, as far as the resolver can tell (but this would be caught later by the verifier) - } - - type Sbyte = x: int | 0 <= x < 256 // supports auto-initialization - type SMyInt = int // supports auto-initialization - type SSixOrMore = x | 6 <= x ghost witness 6 - type SAnotherSixOrMore = s: SSixOrMore | true ghost witness 6 - type SMySixOrMore = x: SMyInt | 6 <= x ghost witness 6 - type SUnclearA = x: int | true ghost witness 0 // see note about for UnclearA - type SUnclearB = x | 6 <= x // see note about for UnclearB - - method M4(a: Sbyte, b: SMyInt, c: SSixOrMore, d: SAnotherSixOrMore, e: SMySixOrMore, - sua: SUnclearA, sub: SUnclearB) { - P(a); - P(b); - P(c); // error: type of argument is expected to support auto-initialization - P(d); // error: type of argument is expected to support auto-initialization - P(e); // error: type of argument is expected to support auto-initialization - P(sua); // error: as far as the resolver can tell, type of argument does not support auto-initialization - P(sub); // fine, as far as the resolver can tell (but this would be caught later by the verifier) - } -} -abstract module ZI_RefinementAbstract { - type A0 - type A1 - type A2 - type A3 - type B0(0) - type B1(0) - type B2(0) - type B3(0) - type C0(00) - type C1(00) - type C2(00) - type C3(00) - - method Delta() -} -module ZI_RefinementConcrete0 refines ZI_RefinementAbstract { - newtype Kuusi = x | 6 <= x witness 6 // supports auto-initialization - newtype Six = x | 6 <= x ghost witness 6 // does not support auto-initialization - newtype Sesis = x | 6 <= x witness * // possibly empty - type A0 = int - type A1 = Kuusi - type A2 = Six - type A3 = Sesis - type B0 = int - type B1 = Kuusi - type B2 = Six // error: RHS is expected to support auto-initialization - type B3 = Sesis // error: RHS is expected to support auto-initialization - type C0 = int - type C1 = Kuusi - type C2 = Six - type C3 = Sesis // error: RHS is expected to be nonempty -} -module ZI_ExportSource { - export - reveals RGB - provides XYZ - datatype RGB = Red | Green | Blue - datatype XYZ = X | Y | Z -} - -module ZI_RefinementConcrete1 refines ZI_RefinementAbstract { - import Z = ZI_ExportSource - - method P(g: G) - method M(m: Z.RGB, n: Z.XYZ) { - P(m); - P(n); // error: Z.XYZ is not known to support auto-initialization - } - - type A0 - type A1(0) // error: not allowed to change auto-initialization setting - type A2(00) // error: not allowed to change nonempty setting - type B0 // error: not allowed to change auto-initialization setting - type B1(0) - type B2(00) // error: not allowed to change auto-initialization setting - type C0 // error: not allowed to change nonempty setting - type C1(0) // error: not allowed to change auto-initialization setting - type C2(00) - - method Delta< - Q, // error: not allowed to change auto-initialization setting - W, - E(0), - R(0)>() // error: not allowed to change auto-initialization setting -} - -// ----- constructor-less classes with need for initialization ----- - -module ConstructorlessClasses { - class C { // error: must have constructor - var x: int - var s: string - var t: set - var u: T - const c: T - } - - method Test() - { - var c := new C>; - print "int: ", c.x, "\n"; - print "string: ", c.s, "\n"; - print "set>: ", c.t, "\n"; - print "real: ", c.u, "\n"; - print "real: ", c.c, "\n"; - } - - codatatype Co = Suc(Co>) // does not know a known compilable value - codatatype Co2 = CoEnd | CoSuc(Co2) - trait Trait { - var co: Co // has no known initializer - } - class Class extends Trait { // error: must have constructor, because of inherited field "co" - } - - class CoClass0 { // error: must have constructor - const co: Co - } - - class CoClass1 { // fine - const co: Co2 := CoEnd - } - - trait CoTrait { - const co: Co2 := CoEnd - } - class CoClass2 extends CoTrait { // fine - } - - iterator Iter(u: T) yields (v: T) - { - } -} - -module GhostWitness { - type BadGhost_EffectlessArrow = f: A -> B - | true - witness (GhostEffectlessArrowWitness) // error: a ghost witness must use the keyword "ghost" - - type GoodGhost_EffectlessArrow = f: A -> B - | true - ghost witness (GhostEffectlessArrowWitness) - - ghost function GhostEffectlessArrowWitness(a: A): B - { - var b: B :| true; b - } -} - -module BigOrdinalRestrictions { // also see BigOrdinalRestrictionsExtremePred below - method Test() { - var st: set; // error: cannot use ORDINAL as type argument - var p: (int, ORDINAL); // error: cannot use ORDINAL as type argument - var o: ORDINAL; // okay - ghost var f := F(o); // error: cannot use ORDINAL as type argument - f := F'(); // error: cannot use ORDINAL as type argument - f := F'<(char,ORDINAL)>(); // error: cannot use ORDINAL as type argument - var lambda := F'; // error: cannot use ORDINAL as type argument - ParameterizedMethod(o); // error: cannot use ORDINAL as type argument - assert forall r: ORDINAL :: P(r); - assert forall r: (ORDINAL, int) :: F(r.1) < 8; - assert exists x: int, r: ORDINAL, y: char :: P(r) && F(x) == F(y); - var s := set r: ORDINAL | r in {}; // error: cannot use ORDINAL as type argument (to set) - var s' := set r: ORDINAL | true :: 'G'; - ghost var m := imap r: ORDINAL :: 10; - var sq := [o, o]; // error: cannot use ORDINAL as type argument (to seq) - var mp0 := map[o := 'G']; // error: cannot use ORDINAL as type argument (to map) - var mp1 := map['G' := o]; // error: cannot use ORDINAL as type argument (to map) - var w := var h: ORDINAL := 100; h + 40; // okay - var w': (int, ORDINAL); // error: cannot use ORDINAL as type argument - var u: ORDINAL :| u == 15; - var ti: ORDINAL :| assume true; - var u': (ORDINAL, int) :| u' == (15, 15); // error (x2): ORDINAL cannot be a type argument - var ti': (ORDINAL, ORDINAL) :| assume true; // error (x4): ORDINAL cannot be a type argument - var lstLocal := var lst: ORDINAL :| lst == 15; lst; - var lstLocal' := var lst: (ORDINAL, int) :| lst == (15, 15); lst.1; // error: ORDINAL cannot be a type argument - if yt: ORDINAL :| yt == 16 { - ghost var pg := P(yt); - } - if { - case zt: ORDINAL :| zt == 180 => - ghost var pg := P(zt); - } - forall om: ORDINAL // allowed - ensures om < om+1 - { - } - var arr := new int[23]; - forall om: ORDINAL | om == 11 // allowed - { - arr[0] := 0; - } - } - ghost function F(g: G): int - ghost function F'(): int - method ParameterizedMethod(g: G) - ghost predicate P(g: ORDINAL) -} - -module IteratorDuplicateParameterNames { - // each of the following once caused a crash in the resolver - iterator MyIterX(u: char) yields (u: char) // error: duplicate name "u" - iterator MyIterY(us: char) yields (u: char) // error: in-effect-duplicate name "us" -} - -module TernaryTypeCheckinngAndInference { - codatatype Stream = Cons(int, Stream) - - method M(k: nat, K: ORDINAL, A: Stream, B: Stream) - requires A == B - { - // all of the following are fine - assert A ==#[k] B; - assert A ==#[K] B; - assert A ==#[3] B; - var b; - assert A ==#[b] B; - } -} - -module DontQualifyWithNonNullTypeWhenYouMeanAClass { - module Z { - class MyClass { - static const g := 100 - } - method M0() { - var x := MyClass?.g; // error: use MyClass, not MyClass? - assert x == 100; - } - method M1() { - var x := MyClass.g; // that's it! - assert x == 100; - } - method P(from: MyClass) returns (to: MyClass?) { - to := from; - } - method Q() { - var x := MyClass; // error: type used as variable - var y := MyClass?; // error: type used as variable - } - } - - module A { - class MyClass { - static const g := 100 - } - } - - module B { - import A - method M0() { - var x := A.MyClass?.g; // error: use MyClass, not MyClass? - assert x == 100; - } - method M1() { - var x := A.MyClass.g; // that's it! - assert x == 100; - } - method P(from: A.MyClass) returns (to: A.MyClass?) { - to := from; - } - method Q() { - var x := A.MyClass; // error: type used as variable - var y := A.MyClass?; // error: type used as variable - } - } -} - -module UninterpretedModuleLevelConst { - type Six = x | 6 <= x witness 6 - type Odd = x | x % 2 == 1 ghost witness 7 - const S: Six // fine - const X: Odd // error: the type of a non-ghost static const must have a known initializer - - class MyClass { } - const Y: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer - ghost const Y': MyClass // error: the type of a ghost static const must be known to be nonempty - - class AnotherClass { // fine, the class itself is not required to have a constructor, because the bad fields are static - static const k := 18 - static const W: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer - static const U: Six // fine, since Six has a non-ghost witness and thus has a known initializer - static const O: Odd // error: the type of a non-ghost static const must have a known initializer - const u: Six - } - - trait Trait { - static const k := 18 - static const W: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer - static const U: Six // fine, since Six has a non-ghost witness and thus has a known initializer - static const O: Odd // error: the type of a non-ghost static const must have a known initializer - } - class ClassyTrait extends Trait { // fine, since the bad fields in Trait are static - } - - trait InstanceConst { - const w: MyClass - } - class Instance extends InstanceConst { // error: because of "w", must declare a constructor - } - - trait GhostTr { - ghost const w: MyClass // the responsibility to initialize "w" lies with any class that implements "GhostTr" - } - class GhostCl extends GhostTr { // error: w and z need initialization, so GhostCl must have a constructor - ghost const z: MyClass - } -} - -module NonThisConstAssignments { - const X: int - - class Cla { - constructor () { - X := 15; // error: can never assign to static const (this used to crash) - Clb.Y := 15; // error: can never assign to static const (this used to crash) - } - } - - class Clb { - static const Y: int - constructor () { - Y := 15; // error: cannot ever assign to a static const (this used to crash) - } - } - - class Clc { - const Z: int - constructor (c: Clc) { - c.Z := 15; // error: can assign to const only for 'this' (this used to crash) - } - } - - class Cld { - const Z: int - constructor () { - Z := 15; - } - } -} - -module ConstGhostRhs { - class S { - const m: int := n // error: use of ghost to assign non-ghost field - ghost const n: int - } - const a: int := b // error: use of ghost to assign non-ghost field - ghost const b: int - - class S' { - ghost const m': int := n' - const n': int - } - ghost const a': int := b' - const b': int - -} - -module Regression15 { - predicate F(i: int, j: int) { true } - function S(i: int): set { {i} } - method M0() returns (b: bool) { - b := forall i, j | j <= i <= 100 && i <= j < 100 :: true; // error: this bogus cyclic dependency was once allowed - } - method M4() returns (b: bool) { - b := forall i, j :: j <= i < 100 && j in S(i) ==> F(i,j); // error: this bogus cyclic dependency was once allowed - } -} - -module AllocDepend0a { - class Class { - const z := if {} == set c: Class | true then 5 else 4 // error: condition depends on alloc (also not compilable, but that's not reported in the same pass) - } - const y := if {} == set c: Class | true then 5 else 4 // error: condition depends on alloc (also not compilable, but that's not reported in the same pass) - newtype byte = x | x < 5 || {} == set c: Class | true // error: condition not allowed to depend on alloc - type small = x | x < 5 || {} == set c: Class | true // error: condition not allowed to depend on alloc -} -module AllocDepend0b { - class Class { } - method M() returns (y: int, z: int) { - z := if {} == set c: Class | true then 5 else 4; // error: not compilable - y := if {} == set c: Class | true then 5 else 4; // error: not compilable - } -} -module AllocDepend1 { - class Class { } - predicate P(x: int) { - x < 5 || {} == set c: Class | true // error: function not allowed to depend on alloc - } -} - -module AllocDepend2 { - class Klass { - const z := if exists k: Klass :: allocated(k) then 3 else 4 // error (x2): condition depends on alloc - } - const y := if exists k: Klass :: allocated(k) then 3 else 4 // error (x2): condition depends on alloc - newtype byte = x | x < 5 || exists k: Klass :: allocated(k) // error: condition not allowed to depend on alloc - type small = x | x < 5 || exists k: Klass :: allocated(k) // error: condition not allowed to depend on alloc -} -module AllocDepend3 { - class Klass { } - predicate P(x: int) { - x < 5 || exists k: Klass :: allocated(k) // error: function not allowed to depend on alloc - } -} - -module AllocDepend4 { - class Xlass { - const z := if var k: Xlass? := null; allocated(k) then 3 else 4 // error (x2): condition depends on alloc - } - const y := if var k: Xlass? := null; allocated(k) then 3 else 4 // error (x2): condition depends on alloc - newtype byte = x | x < 5 || var k: Xlass? := null; allocated(k) // error: condition not allowed to depend on alloc - type small = x | x < 5 || var k: Xlass? := null; allocated(k) // error: condition not allowed to depend on alloc -} -module AllocDepend5 { - class Xlass { } - predicate P(x: int) { - x < 5 || var k: Xlass? := null; allocated(k) // error: function not allowed to depend on alloc - } -} - -module AbstemiousCompliance { - codatatype EnormousTree = Node(left: EnormousTree, val: X, right: EnormousTree) - - ghost function {:abstemious} Five(): int { - 5 // error: an abstemious function must return with a co-constructor - } - - ghost function {:abstemious} Id(t: EnormousTree): EnormousTree - { - t // to be abstemious, a parameter is as good as a co-constructor - } - - ghost function {:abstemious} IdGood(t: EnormousTree): EnormousTree - { - match t - case Node(l, x, r) => Node(l, x, r) - } - - ghost function {:abstemious} AlsoGood(t: EnormousTree): EnormousTree - { - Node(t.left, t.val, t.right) - } - - ghost function {:abstemious} UniformTree(x: X): EnormousTree - { - Node(UniformTree(x), x, UniformTree(x)) - } - - ghost function {:abstemious} AlternatingTree(x: X, y: X): EnormousTree - { - Node(AlternatingTree(y, x), x, AlternatingTree(y, x)) - } - - ghost function {:abstemious} AnotherAlternatingTree(x: X, y: X): EnormousTree - { - var t := Node(AlternatingTree(x, y), y, AlternatingTree(x, y)); - Node(t, x, t) - } - - ghost function {:abstemious} NonObvious(x: X): EnormousTree - { - AlternatingTree(x, x) // error: does not return with a co-constructor - } - - ghost function {:abstemious} BadDestruct(t: EnormousTree): EnormousTree - { - Node(t.left, t.val, t.right.right) // error: cannot destruct t.right - } - - ghost function {:abstemious} BadMatch(t: EnormousTree): EnormousTree - { - match t.right // error: cannot destruct t.right - case Node(a, x, b) => - Node(a, x, b) - } - - ghost function {:abstemious} BadEquality(t: EnormousTree, u: EnormousTree, v: EnormousTree): EnormousTree - { - if t == u then // error: cannot co-compare - Node(t.left, t.val, t.right) - else if u != v then // error: cannot co-compare - Node(u.left, u.val, u.right) - else - Node(v.left, v.val, v.right) - } - - ghost function {:abstemious} Select(b: bool, t: EnormousTree, u: EnormousTree): EnormousTree - { - if b then t else u // abstemious yes: parameters are as good as a co-constructors - } - - ghost function {:abstemious} Select'(b: bool, t: EnormousTree, u: EnormousTree): EnormousTree - { - if b then - Node(t.left, t.val, t.right) // fine, too - else - Node(u.left, u.val, u.right) // fine, too - } -} - -module BigOrdinalRestrictionsExtremePred { - least predicate Test() { - var st: set := {}; // error: cannot use ORDINAL as type argument - var p: (int, ORDINAL) := (0,0); // error: cannot use ORDINAL as type argument - var o: ORDINAL := 0; // okay - ghost var f := F(o); // error: cannot use ORDINAL as type argument - var f := F'(); // error: cannot use ORDINAL as type argument - var f := F'<(char,ORDINAL)>(); // error: cannot use ORDINAL as type argument - var lambda := F'; // error: cannot use ORDINAL as type argument - ParameterizedLemma(o); // error: cannot use ORDINAL as type argument - assert forall r: ORDINAL :: P(r); // error: cannot quantify over ORDINAL here - assert forall r: (ORDINAL, int) :: F(r.1) < 8; // error: cannot quantify over ORDINAL here - assert exists x: int, r: ORDINAL, y: char :: P(r) && F(x) == F(y); // error: cannot quantify over ORDINAL here - var s := set r: ORDINAL | r in {}; // error (x2): cannot use ORDINAL as type argument (to set) - var s' := set r: ORDINAL | true :: 'G'; // error: cannot use ORDINAL here - ghost var m := imap r: ORDINAL :: 10; // error (x2): cannot use ORDINAL here - var sq := [o, o]; // error: cannot use ORDINAL as type argument (to seq) - var mp0 := map[o := 'G']; // error: cannot use ORDINAL as type argument (to map) - var mp1 := map['G' := o]; // error: cannot use ORDINAL as type argument (to map) - var w := var h: ORDINAL := 100; h + 40; // okay - var w': (int, ORDINAL) := (8,8); // error: cannot use ORDINAL as type argument - var u: ORDINAL :| u == 15; - var ti: ORDINAL :| true; - var u': (ORDINAL, int) :| u' == (15, 15); // error (x2): ORDINAL cannot be a type argument - var ti': (ORDINAL, ORDINAL) :| true; // error (x2): ORDINAL cannot be a type argument - var lstLocal := var lst: ORDINAL :| lst == 15; lst; - var lstLocal' := var lst: (ORDINAL, int) :| lst == (15, 15); lst.1; // error: ORDINAL cannot be a type argument - var gr := if yt: ORDINAL :| yt == 16 then - ghost var pg := P(yt); 5 - else - 7; - calc { - 100; - == { - forall om: ORDINAL // allowed - ensures om < om+1 - { - } - } - 100; - } - true - } - ghost function F(g: G): int - ghost function F'(): int - lemma ParameterizedLemma(g: G) - ghost predicate P(g: ORDINAL) -} - -// ----- label domination ----- - -module LabelDomination { - method DuplicateLabels(n: int) { - var x; - if (n < 7) { - label L: x := x + 1; - } else { - label L: x := x + 1; - } - assert old@L(true); // error: L is not available here - label L: x := x + 1; - assert old@L(true); - label L: x := x + 1; // error: duplicate label - assert old@L(true); - - { - label K: - x := x + 1; - } - assert old@K(true); - label K: // error: duplicate label - assert old@K(true); - } - - datatype Color = A | B | C - method Branches(n: int, c: Color) { - var x: int; - if n < 2 { - label X: x := x + 1; - } else if n < 4 { - label X: x := x + 1; - } else { - label X: x := x + 1; - } - if * { - label X: x := x + 1; - } else { - label X: x := x + 1; - } - - if { - case true => - label X: x := x + 1; - case true => - label X: x := x + 1; - } - - var i := 0; - while i < x { - label X: x := x + 1; - i := i + 1; - } - - i := 0; - while { - case i < x => - label X: x := x + 1; - i := i + 1; - case i < x => - label X: x := x + 1; - i := i + 1; - } - - match c { - case A => - label X: x := x + 1; - case B => - label X: x := x + 1; - case C => - label X: x := x + 1; - } - - label X: x := x + 1; // all okay - } - - method A0() { - label Q: - assert true; - } - method A1() { - label Q: - assert true; - } - - class MyClass { - var x: int - method LabelNotInScope_Old(y: int) { - label GoodLabel: - if y < 5 { - label Treasure: - assert true; - } - assert old(x) == old@Treasure(x); // error: no label Treasure in scope - assert 10 == old@WonderfulLabel(x); // error: no label WonderfulLabel in scope - assert old(x) == old@GoodLabel(x); - assert old(x) == old@FutureLabel(x); // error: no label FutureLabel in scope - label FutureLabel: {} - } - method LabelNotInScope_Unchanged(y: int) { - label GoodLabel: - if y < 5 { - label Treasure: - assert true; - } - assert unchanged@Treasure(`x); // error: no label Treasure in scope - assert unchanged@WonderfulLabel(this); // error: no label WonderfulLabel in scope - assert unchanged@GoodLabel(this); - assert unchanged@FutureLabel(this); // error: no label FutureLabel in scope - label FutureLabel: {} - } - method LabelNotInScope_Fresh(y: int, c: MyClass) { - label GoodLabel: - if y < 5 { - label Treasure: - assert true; - } - assert fresh@Treasure(c); // error: no label Treasure in scope - assert fresh@WonderfulLabel(c); // error: no label WonderfulLabel in scope - assert fresh@GoodLabel(c); - assert fresh@FutureLabel(c); // error: no label FutureLabel in scope - label FutureLabel: {} - } - } -} - -// ----- bad use of types without auto-initializers ----- - -module Initialization { - datatype Yt = MakeYt(x: int, y: Y) - type Even = x | x % 2 == 0 - type Odd = x | x % 2 == 1 witness 17 - type GW = x | x % 2 == 1 ghost witness 17 - method DefiniteAssignmentViolation() returns (e: Yt, o: Yt, g: Yt) - { - } // no resolution errors (but verification errors, see NonZeroInitialization.dfy) - method ArrayElementInitViolation() returns (e: array>, o: array>, g: array>) - { - e := new Yt[20]; - o := new Yt[20]; - g := new Yt[20]; - } // no resolution errors (but verification errors, see NonZeroInitialization.dfy) - method GimmieOne() returns (g: G) - { - } - method TypeParamViolation() returns (e: Yt, o: Yt, g: Yt) - { - e := GimmieOne>(); - o := GimmieOne>(); - g := GimmieOne>(); // error: cannot pass Yt to a (0)-parameter - } -} - -// ----------------- regression tests ---------------------------------------------------- - -module FreshTypeInferenceRegression { - class MyClass { - method M(N: nat) - { - var i, os := 0, {}; - while i < N - invariant fresh(os) - invariant forall o :: o in os ==> fresh(o.inner) // error: type of "o" not yet known (this once caused a crash) - { - var o := new Outer(); - os, i := os + {o}, i + 1; - } - } - } - - class Outer { - const inner: Inner - constructor () - ensures fresh(inner) - { - inner := new Inner(); - } - } - - class Inner { - constructor () - } -} - -module RegressionTest { - class Cache { - method Lookup(K: X) returns (V: X) - { - V := Cache[K]; // error: Cache is not a field but a type - } - } -} - -module ExistsImpliesWarning { - method M(a: array, b: array) - requires a.Length == b.Length - requires exists i :: 0 <= i < a.Length ==> a[i] == b[i] // warning - requires exists i :: true && (0 <= i < a.Length ==> a[i] == b[i]) - requires exists i :: (0 <= i < a.Length ==> a[i] == b[i]) - requires exists i | 0 <= i < a.Length :: true ==> a[i] == b[i] - requires exists i | 0 <= i < a.Length :: a[i] == b[i] - requires exists i | 0 <= i < a.Length ==> a[i] == b[i] :: true - requires exists i :: !(0 <= i < a.Length) || a[i] == b[i] - requires exists i :: a[i] == b[i] <== 0 <= i < a.Length // warning - requires exists i :: !(0 <= i < a.Length && a[i] != b[i]) - requires exists i :: 0 <= i < a.Length && a[i] == b[i] - requires exists i :: true ==> (0 <= i < a.Length && a[i] == b[i]) // warning - { - } - - method N(a: array, b: array) - requires a.Length == b.Length - requires forall i :: 0 <= i < a.Length ==> a[i] == b[i] - requires forall i :: true && (0 <= i < a.Length ==> a[i] == b[i]) - requires forall i :: (0 <= i < a.Length ==> a[i] == b[i]) - requires forall i | 0 <= i < a.Length :: true ==> a[i] == b[i] - requires forall i | 0 <= i < a.Length :: a[i] == b[i] - requires forall i | 0 <= i < a.Length ==> a[i] == b[i] :: true - requires forall i :: !(0 <= i < a.Length) || a[i] == b[i] - requires forall i :: a[i] == b[i] <== 0 <= i < a.Length - requires forall i :: !(0 <= i < a.Length && a[i] != b[i]) - requires forall i :: 0 <= i < a.Length && a[i] == b[i] - requires forall i :: true ==> (0 <= i < a.Length && a[i] == b[i]) - { - } -} - -// --------------- ghost (regression) tests, receivers ------------------------------------- - -module GhostReceiverTests { - class C { - ghost function F(x: int): int { 3 } - function G(x: int): int { 4 } - lemma L(x: int) { } - method M(x: int) { } - } - method Caller(x: int, ghost z: int, c: C, ghost g: C) { - { - var y; - y := c.F(x); // error: LHS is non-ghost, so RHS cannot use ghost function F - y := g.F(x); // error: LHS is non-ghost, so RHS cannot use ghost function F - y := c.G(x); - y := g.G(x); // error: LHS is non-ghost, so RHS cannot use ghost variable g - } - { - // all of the these are fine, because: the LHS is ghost and, therefore, the whole statement is - ghost var y; - y := c.F(x); - y := g.F(x); - y := c.G(x); - y := g.G(x); - } - { - // all of the these are fine, because: the LHS is ghost and, therefore, the whole statement is - ghost var y; - y := c.F(z); - y := g.F(z); - y := c.G(z); - y := g.G(z); - } - c.L(x); - g.L(x); - c.M(x); - g.M(x); // error: cannot pass ghost receiver to compiled method - } -} - -// --------------- ghost RHS of constants (regression) tests ------------------------------------- - -module GhostRhsConst { - class C { - ghost function F(n: nat): nat { n } // a ghost function - static ghost function G(n: nat): nat { n } // a ghost function - const b := F(0) // error: RHS uses a ghost function - static const u := G(0) // error: RHS uses a ghost function - } - - trait R { - ghost function F(n: nat): nat { n } // a ghost function - static ghost function G(n: nat): nat { n } // a ghost function - const b := F(0) // error: RHS uses a ghost function - static const u := G(0) // error: RHS uses a ghost function - } -} - -// --------------- errors from nested modules ------------------------------------- - -module ErrorsFromNestedModules { - method M() { - U.V.Test(); - UU.V.Test(); - } - - module U { // regression test: since U is rather empty, this had once caused an error - module V { - method Test() { - } - module W { - const x1 := 12 * false // error: bad types - } - } - } - - module UU.V { // same regression as above - method Test() { - } - module W { - const x1 := 12 * false // error: bad types - } - } -} - -// --------------- name clashes related to prefix-named modules ------------------------------------- - -module NameClashes { - module U.G { - } - module U { - class G { } // error: duplicate name: G - class H { } - } - module U.H { // error: duplicate name: H - } -} - -// --------------- regression ghost tests ------------------------------------- - -module RegressionGhostTests { - class Cell { - var data: int - } - - method field(x: Cell) - modifies x - { - ghost var y := x; - x.data := 42; - y.data := 42; // error: assignment to non-ghost field depends on a ghost - (assert x == y; x).data := 42; - (assert x == y; y).data := 42; // error: assignment to non-ghost field depends on a ghost - } - - method arr(a: array) - requires 5 < a.Length - modifies a - { - ghost var b := a; - ghost var i := 5; - a[i] := 42; // error: assignment to non-ghost field depends on a ghost - b[5] := 42; // error: assignment to non-ghost field depends on a ghost - } - - method arr2(a: array2) - requires 5 < a.Length0 && 5 < a.Length1 - modifies a - { - ghost var b := a; - ghost var i := 5; - a[i,5] := 42; // error: assignment to non-ghost field depends on a ghost - a[5,i] := 42; // error: assignment to non-ghost field depends on a ghost - b[5,5] := 42; // error: assignment to non-ghost field depends on a ghost - } -} - -// --------------- regression test const in frame expression ------------------------------ - -module RegressionConstFrameExpression { - class C { - const x: int - var y: int - } - method m(c: C) - modifies c`x - modifies c`y - ensures unchanged(c`x) - ensures unchanged(c) - { - - } -} - -// --------------- change in language semantics to forbid !! on maps ------------------------------ - -module MapDisjointnessNoMore { - method M(a: map, b: map) { - assert a !! b; // error: !! is (no longer) support on maps - assert a.Keys !! b.Keys; // instead, this is the way to do it - } -} - -// --------------- expect statements ------------------------------ - -module ExpectStatements { - - ghost function UnsafeDivide(a: int, b: int): int { - expect b != 0; // expect statement is not allowed in this context - a / b - } - - method M() { - ghost var g := 5; - expect forall i : int :: i == i; // error: quantifiers in non-ghost contexts must be compilable - expect false, if g == 5 then "boom" else "splat"; // error: ghost variables are allowed only in specification contexts - } -} - -// --------------- type-parameter scopes ------------------------------ - -module TypeParameterScopes { - class C { - function G(): X - method M(f: X) { - var h: X := f; - var k: X := G(); // error: this is the wrong X - } - function F(f: X): int { - var h: X := f; - var k: X := G(); // error: this is the wrong X - 10 - } - } -} - -// --------------- type of function members (regression tests) ------------------------------ - -module TypeOfFunctionMember { - ghost function Fo(x: X): int - - lemma M() { - // Both of the following once crashed the type checker - var rd := Fo.reads; - var rq := Fo.requires; - } -} - -// --------------- update operations ------------------------------ - -module CollectionUpdates { - // Update operations on collections must have the right types, modulo subset types. - // For verification errors, see Maps.dfy. - trait Trait { } - class Elem extends Trait { } - - method UpdateValiditySeq(d: Trait, e: Elem) { - var s: seq := [e, e, e, e, e]; - s := s[1 := d]; // error: d is not an Elem (and is not a subset type of it, either) - } - method UpdateValidityMultiset(d: Trait) { - var s: multiset; - s := s[d := 5]; // error: element value is not a Elem - } - method UpdateValidityMap(d: Trait, e: Elem) { - var m: map; - if * { - m := m[d := e]; // error: key is not a Elem - } else { - m := m[e := d]; // error: value is not a Elem - } - } -} - -// --------------- update operations ------------------------------ - -module MoreAutoInitAndNonempty { - type A(0) - type B(00) - type C - - method Q(f: F) - method P(g: G) - method R(h: H) - - function FQ(f: F): int - function FP(g: G): int - function FR(h: H): int - - method M(x: X, y: Y, z: Z) - { - Q(x); - P(x); - R(x); - Q(y); // error: auto-init mismatch - P(y); - R(y); - Q(z); // error: auto-init mismatch - P(z); // error: auto-init mismatch - R(z); - } - - method N(x: X, y: Y, z: Z) returns (u: int) - { - u := FQ(x); - u := FP(x); - u := FR(x); - u := FQ(y); // error: auto-init mismatch - u := FP(y); - u := FR(y); - u := FQ(z); // error: auto-init mismatch - u := FP(z); // error: auto-init mismatch - u := FR(z); - } -} - -// --------------- ghost function error messages ------------------------------ - -module GhostFunctionErrorMessages { - ghost function GhostFunction(): int - ghost predicate GhostPredicate() - least predicate LeastPredicate() - greatest predicate GreatestPredicate() - twostate function TwoFunction(): int - twostate predicate TwoPredicate() - - method GhostsUsedInCompiledContexts() { - var x, b; - x := GhostFunction(); // error - b := GhostPredicate(); // error - b := LeastPredicate(); // error - b := GreatestPredicate(); // error - x := TwoFunction(); // error - b := TwoPredicate(); // error - } -} - -module TypeParameterCount { - ghost function F0(): int - ghost function F1(): int - ghost function F2(): int - method M0() - method M1() - method M2() - - method TestFunction() { - var x; - x := F0(); - x := F1(); // type argument inferred - x := F2(); // type arguments inferred - x := F0(); // error: wrong number of type parameters - x := F1(); - x := F2(); // error: wrong number of type parameters - x := F0(); // error: wrong number of type parameters - x := F1(); // error: wrong number of type parameters - x := F2(); - } - - method TestMethods() { - M0(); - M1(); // type argument inferred - M2(); // type arguments inferred - M0(); // error: wrong number of type parameters - M1(); - M2(); // error: wrong number of type parameters - M0(); // error: wrong number of type parameters - M1(); // error: wrong number of type parameters - M2(); - } -} - -module AutoGhostRegressions { - datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) - - method Test() { - var q := Quad(20, 30, 40, 50); - print q, "\n"; - - var Quad(a, b, c, d) := q; - print c, "\n"; // error: c is ghost (this was once not handled correctly) - - match q { - case Quad(r, s, t, u) => - print t, "\n"; // error: t is ghost - } - - ghost var p := Quad(20, 30, 40, 50); - var Quad(a', b', c', d') := p; - print a', "\n"; // error: a' is ghost - print c', "\n"; // error: c' is ghost - } - - datatype NoEquality = NoEquality(ghost u: int) - newtype NT = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context - type ST = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context -} - -module TypeCharacteristicsInGhostCode { - function MustBeNonempty(): int { 5 } - function MustBeAutoInit(): int { 5 } - function MustSupportEquality(): int { 5 } - function NoReferences(): int { 5 } - - type PossiblyEmpty = x: int | true witness * - type Nonempty = x: int | true ghost witness 0 - datatype NoEquality = NoEquality(ghost u: int) - class Class { } - type Good = bool - - method TestCompiled() - { - var w; - - w := MustBeNonempty(); // error - w := MustBeNonempty(); - w := MustBeNonempty(); - w := MustBeNonempty(); - w := MustBeNonempty(); - w := MustBeNonempty(); // error (a hint is given) - - w := MustBeAutoInit(); // error - w := MustBeAutoInit(); // error - w := MustBeAutoInit(); - w := MustBeAutoInit(); - w := MustBeAutoInit(); - w := MustBeAutoInit(); // error (a hint is given) - - w := MustSupportEquality(); - w := MustSupportEquality(); - w := MustSupportEquality(); // error - w := MustSupportEquality(); - w := MustSupportEquality(); - w := MustSupportEquality(); // error (a hint is given) - - w := NoReferences(); - w := NoReferences(); - w := NoReferences(); - w := NoReferences(); // error - w := NoReferences(); - w := NoReferences(); // error (a hint is given) - } - - method TestGhost() - { - ghost var w; - - w := MustBeNonempty(); // error - w := MustBeNonempty(); - w := MustBeNonempty(); - w := MustBeNonempty(); - w := MustBeNonempty(); - - w := MustBeAutoInit(); // error - w := MustBeAutoInit(); // fine, because the call appears in a ghost context - w := MustBeAutoInit(); - w := MustBeAutoInit(); - w := MustBeAutoInit(); - - w := MustSupportEquality(); - w := MustSupportEquality(); - w := MustSupportEquality(); - w := MustSupportEquality(); - w := MustSupportEquality(); - - w := NoReferences(); - w := NoReferences(); - w := NoReferences(); - w := NoReferences(); // error - w := NoReferences(); - } - - function FF(a: bool, ghost b: bool): int { 5 } - method MM(a: bool, ghost b: bool) { } - function GetInt(): int { 2 } - method GhostContexts(x: T, y: T) { - var r; - r := FF(x == y, true); // error: T must support equality - r := FF(true, x == y); // no problem, since this is a ghost context - MM(x == y, true); // error: T must support equality - MM(true, x == y); // no problem, since this is a ghost context - - r := FF(GetInt() == 0, true); // error: type argument must support equality - r := FF(true, GetInt() == 0); // okay, since this is a ghost context - MM(GetInt() == 0, true); // error: type argument must support equality - MM(true, GetInt() == 0); // okay, since this is a ghost context - - var q0 := Quad(GetInt() == 0, true, true, true); // error: type argument must support equality - var q1 := Quad(true, true, GetInt() == 0, true); // fine, since in a ghost context - } - - datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) - datatype QuadEq = QuadEq(0: T, 1: T, ghost 2: U, ghost 3: U) - - method VarDecls(x: T, y: T) { - var a := x == y; // error: this requires T to support equality - ghost var b := x == y; // fine - - var q := Quad(x, y, x, y); - var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost - var c := k == l; // error: this requires T to support equality - var d := m == n; // fine, since d is implicitly ghost - d, c, d := m == n, k == l, m == n; // error: "k == l" requires T to support equality - - var q' := QuadEq([x], [x], [0], [0]); // error: seq requires T to support equality - var q'' := QuadEq([0], [0], [x], [x]); // error: seq requires T to support equality - } - - newtype NT = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context - type ST = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context - - method LetVarDecls(x: T, y: T) { - var lhs; - lhs := - var a := x == y; // error: this requires T to support equality - 0; - lhs := - ghost var b := x == y; // fine - 0; - - var q := Quad(x, y, x, y); - var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost - lhs := - var c := k == l; // error: this requires T to support equality - 0; - lhs := - var d := m == n; // fine, since d is implicitly ghost - 0; - - ghost var ghostLhs; - ghostLhs := - var a := x == y; // fine - 0; - ghostLhs := - ghost var b := x == y; // fine - 0; - } - - datatype DatatypeHasMembersToo = Abc | Def { - method Test() { - var w; - w := MustBeNonempty(); // error - w := MustBeAutoInit(); // error - w := MustBeAutoInit(); // error - w := MustSupportEquality(); // error - w := NoReferences(); // error - } - } - - newtype NewtypeHasMembersToo = x: int | x == MustBeNonempty() // error: constraint has bad type instantiation - witness MustBeNonempty() // error: witness expression has bad type instantiation - { - method Test() { - var w; - w := MustBeNonempty(); // error - w := MustBeAutoInit(); // error - w := MustBeAutoInit(); // error - w := MustSupportEquality(); // error - w := NoReferences(); // error - } - } - - type SubsetTypeHasExpressionToo = x: int | x == MustBeNonempty() // error: constraint has bad type instantiation - witness MustBeNonempty() // error: witness expression has bad type instantiation - - newtype NT_CompiledWitness = x | 0 <= x - witness MustSupportEquality() // error - newtype NT_GhostWitness = x | 0 <= x - ghost witness MustSupportEquality() // fine, since it's ghost - type ST_CompiledWitness = x | 0 <= x - witness MustSupportEquality() // error - type ST_GhostWitness = x | 0 <= x - ghost witness MustSupportEquality() // fine, since it's ghost - - trait - {:MyAttribute MustSupportEquality(), MustBeNonempty()} // error: about MustBeNonempty (no prob with (==)) - MyTrait - { - const x := (CallMyLemma(MustSupportEquality(), MustBeNonempty()); 23) // error: about MustBeNonempty (no prob with (==)) - } - lemma CallMyLemma(x: int, y: int) -} - -module MoreAutoGhostTests { - datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) - - method SomeLetVarsAreGhost0(q: Quad) returns (r: int) { - r := - var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost - k + l; - } - - method SomeLetVarsAreGhost1(q: Quad) returns (r: int) { - r := - var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost - m; // error: m is ghost - } - - method AssignSuchThat(ghost m: int) { - var d :| d == m; // error: LHS is not inferred to be ghost for :| - } - - function LetSuchThat(ghost m: int): int { - var d :| d == m; // error: LHS is not inferred to be ghost for :| - 0 - } -} - -module RelaxedAutoInitChecking { - // In a ghost context, the (==) is never relevant. Therefore, checking of adherence to (==) for - // type arguments is suppressed in ghost contexts. - // Similarly, in a ghost context, there's no difference between (0) and (00). Therefore, a - // formal parameter that expects (0) can take either a (0) or a (00) in a ghost context. - - function MustBeNonempty(): int { 5 } - function MustBeAutoInit(): int { 5 } - function MustSupportEquality(): int { 5 } - function NoReferences(): int { 5 } - - type PossiblyEmpty = x: int | true witness * - type Nonempty = x: int | true ghost witness 0 - datatype NoEquality = NoEquality(ghost u: int) - class Class { } - type Good = bool - - method M(compiledValue: int, ghost ghostValue: int) - - method TestCompiled() - { - M(MustBeNonempty(), 0); // error - M(MustBeNonempty(), 0); - M(MustBeNonempty(), 0); - M(MustBeNonempty(), 0); - M(MustBeNonempty(), 0); - - M(MustBeAutoInit(), 0); // error - M(MustBeAutoInit(), 0); // error - M(MustBeAutoInit(), 0); - M(MustBeAutoInit(), 0); - M(MustBeAutoInit(), 0); - - M(MustSupportEquality(), 0); - M(MustSupportEquality(), 0); - M(MustSupportEquality(), 0); // error - M(MustSupportEquality(), 0); - M(MustSupportEquality(), 0); - - M(NoReferences(), 0); - M(NoReferences(), 0); - M(NoReferences(), 0); - M(NoReferences(), 0); // error - M(NoReferences(), 0); - } - - method TestGhost() - { - M(0, MustBeNonempty()); // error - M(0, MustBeNonempty()); - M(0, MustBeNonempty()); - M(0, MustBeNonempty()); - M(0, MustBeNonempty()); - - M(0, MustBeAutoInit()); // error - M(0, MustBeAutoInit()); // this is fine in a ghost context - M(0, MustBeAutoInit()); - M(0, MustBeAutoInit()); - M(0, MustBeAutoInit()); - - M(0, MustSupportEquality()); - M(0, MustSupportEquality()); - M(0, MustSupportEquality()); // this is fine in a ghost context - M(0, MustSupportEquality()); - M(0, MustSupportEquality()); - - M(0, NoReferences()); - M(0, NoReferences()); - M(0, NoReferences()); - M(0, NoReferences()); // error - M(0, NoReferences()); - } -} - -// --------------- let-such-that ghost regressions ------------------------------ - -module LetSuchThatGhost { - ghost predicate True(t: T) { true } - - function F(s: set): int - requires s != {} - { - // once, the RHS for p was (bogusly) considered ghost, which made p ghost, - // which made this body illegal - var p := - var e :| e in s; - true; - if p then 6 else 8 - } - - function G(s: set): int - requires s != {} - { - // again, e and p are both non-ghost - var p := - var e :| e in s; - e == e; - if p then 6 else 8 - } - - function H(s: set): int - requires s != {} - { - // here, e is ghost, but p is still not - var p := - var e :| e in s && True(e); - true; - if p then 6 else 8 - } - - function I(s: set): int - requires s != {} - { - // here, e is ghost, and therefore so is p - var p := - var e :| e in s && True(e); - e == e; - if p then 6 else 8 // error: p is ghost - } -} - -// --------------- hint restrictions in non-while loops ------------------------------ - -module HintRestrictionsOtherLoops { - class C { - ghost function F(): int - { - calc { - 6; - { var x := 8; - while - modifies this // error: cannot use a modifies clause on a loop inside a hint - { - case x != 0 => x := x - 1; - } - } - 6; - { for i := 0 to 8 - modifies this // error: cannot use a modifies clause on a loop inside a hint - { - } - } - 6; - } - 5 - } - } -} - -// --------------- regressions: types of arguments to fresh, unchanged, modifies, reads ------------------------------ - -module FrameTypes { - // ----- fresh takes an expression as its argument - - method FreshArgumentType0( - o: object, - s: set, ss: iset, - q: seq, - ms: multiset, - mp: map, mp2: map, - im: imap) - { - ghost var b; - b := fresh(o); - b := fresh(s); - b := fresh(ss); - b := fresh(q); - b := fresh(ms); // error: wrong argument type for fresh - b := fresh(mp); // error: wrong argument type for fresh - b := fresh(mp2); // error: wrong argument type for fresh - b := fresh(im); // error: wrong argument type for fresh - } - - method FreshArgumentType1(x: int, s: set, ss: iset, q: seq) { - ghost var b; - b := fresh(x); // error: wrong argument type for fresh - b := fresh(s); // error: wrong argument type for fresh - b := fresh(ss); // error: wrong argument type for fresh - b := fresh(q); // error: wrong argument type for fresh - } - - method FreshArgumentType2(f: int -> int, g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int) { - ghost var b; - b := fresh(f); // error: wrong argument type for fresh - b := fresh(g); // error: wrong argument type for fresh - b := fresh(h); // error: wrong argument type for fresh - b := fresh(j); // error: wrong argument type for fresh - b := fresh(k); // error: wrong argument type for fresh - } - - // ----- unchanged, modifies, and reads take frame expressions as their arguments - - method UnchangedArgumentType0( - o: object, - s: set, ss: iset, - q: seq, - ms: multiset, - mp: map, mp2: map, - im: imap) - { - ghost var b; - b := unchanged(o); - b := unchanged(s); - b := unchanged(ss); - b := unchanged(q); - b := unchanged(ms); - b := unchanged(mp); // error: wrong argument type for unchanged - b := unchanged(mp2); // error: wrong argument type for unchanged - b := unchanged(im); // error: wrong argument type for unchanged - } - - method UnchangedArgumentType1(x: int, s: set, ss: iset, q: seq) { - ghost var b; - b := unchanged(x); // error: wrong argument type for unchanged - b := unchanged(s); // error: wrong argument type for unchanged - b := unchanged(ss); // error: wrong argument type for unchanged - b := unchanged(q); // error: wrong argument type for unchanged - } - - method UnchangedArgumentType2( - f: int -> int, - g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, - l: bool -> multiset, m: bool -> map) - { - ghost var b; - b := unchanged(f); // error: wrong argument type for unchanged - b := unchanged(g); // error: wrong argument type for unchanged - b := unchanged(h); // error: wrong argument type for unchanged - b := unchanged(i); // error: wrong argument type for unchanged - b := unchanged(j); // error: wrong argument type for unchanged - b := unchanged(k); // error: wrong argument type for unchanged - b := unchanged(l); // error: wrong argument type for unchanged - b := unchanged(m); // error: wrong argument type for unchanged - } - - method ModifiesArgumentType0( - o: object, - s: set, ss: iset, - q: seq, - ms: multiset, - mp: map, mp2: map, - im: imap) - modifies o - modifies s - modifies ss - modifies q - modifies ms - modifies mp // error: wrong argument type for modifies - modifies mp2 // error: wrong argument type for modifies - modifies im // error: wrong argument type for modifies - { - } - - method ModifiesArgumentType1(x: int, s: set, ss: iset, q: seq) - modifies x // error: wrong argument type for modifies - modifies s // error: wrong argument type for modifies - modifies ss // error: wrong argument type for modifies - modifies q // error: wrong argument type for modifies - { - } - - method ModifiesArgumentType2( - f: int -> int, - g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, - l: bool -> multiset, m: bool -> map) - modifies f // error: wrong argument type for modifies - modifies g // error: wrong argument type for modifies - modifies h // error: wrong argument type for modifies - modifies i // error: wrong argument type for modifies - modifies j // error: wrong argument type for modifies - modifies k // error: wrong argument type for modifies - modifies l // error: wrong argument type for modifies - modifies m // error: wrong argument type for modifies - { - } - - ghost predicate ReadsArgumentType0( - o: object, - s: set, ss: iset, - q: seq, - ms: multiset, - mp: map, mp2: map, - im: imap) - reads o - reads s - reads ss - reads q - reads ms - reads mp // error: wrong argument type for reads - reads mp2 // error: wrong argument type for reads - reads im // error: wrong argument type for reads - { - true - } - - ghost predicate ReadsArgumentType1(x: int, s: set, ss: iset, q: seq) - reads x // error: wrong argument type for reads - reads s // error: wrong argument type for reads - reads ss // error: wrong argument type for reads - reads q // error: wrong argument type for reads - { - true - } - - ghost predicate ReadsArgumentType2( - f: int -> int, - g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, - l: bool -> multiset, m: bool -> map) - reads f // error: wrong argument type for reads - reads g // error: a function must be to a collection of references - reads h - reads i - reads j - reads k // error: wrong argument type for reads - reads l - reads m // error: wrong argument type for reads - { - true - } -} - -module Continue0 { - method BadTargetsLevels(a: int, b: int, c: int) { - for i := 0 to 100 { - for j := 0 to 100 { - for k := 0 to 100 { - if - case k == a => - continue; - case k == b => - break continue; - case k == c => - break break continue; - case k == a + b + c => - break break break continue; // error: too many levels - } - } - } - } - - method BadTargetsLabels(a: int, b: int, c: int) { - label A: - for i := 0 to 100 { - label B0: label B1: - for j := 0 to 100 { - label C: - for k := 0 to 100 { - if - case k == a => - continue C; - case k == b => - continue B0; - case k == b => - continue B1; - case k == c => - continue A; - } - } - } - } - - method NonLoopLabels(a: int, b: int, c: int) { - // the following labels are attached to BlockStmt's, not loops - label X: { - for i := 0 to 100 { - label Y0: label Y1: { - for j := 0 to 100 { - label Z: { - for k := 0 to 100 { - if - case k == a => - continue X; // error: X is not a loop label - case k == b => - continue Y0; // error: Y0 is not a loop label - case k == b => - continue Y1; // error: Y1 is not a loop label - case k == c => - continue Z; // error: Z is not a loop label - } - } - } - } - } - } - } - - method SimpleBadJumps0() { - break; // error: cannot "break" from here - } - - method SimpleBadJumps1() { - continue; // error: cannot "continue" from here - } - - method SimpleBadJumps2() { - label X: { - if - case true => break; // error: cannot "break" from here - case true => continue; // error: cannot "continue" from here - case true => break X; - case true => continue X; // error: X is not a loop label - } - } - - method GhostContinueAssertBy(ghost t: int, ghost u: nat) - { - label L: - for i := 0 to 100 { - assert true by { - for j := 0 to 100 { - if j == t { - break; - } else if j == u { - continue; - } - } - if - case true => break; // error: cannot jump outside the assert-by - case true => continue; // error: cannot jump outside the assert-by - case true => break L; // error: cannot jump outside the assert-by - case true => continue L; // error: cannot jump outside the assert-by - } - } - } -} - -module Continue1 { - method GhostContinueLevels(ghost t: int, ghost u: nat) - { - var m := 0; - for i := 0 to 100 { - if i == t { - // The following "continue" would pass the increment to m - continue; // error: continue from ghost context must target a ghost loop - } - m := m + 1; - } - - for i := 0 to 100 { - m := m + 1; - // The following "break" would potentially pass both increments to m - if i == t { - break; // error: break from ghost context must target a ghost loop - } - m := m + 1; - } - - for i := 0 to 100 { - if i == t { - // Even though there's no statement in the loop body after this ghost if, the continue violates the rule - continue; // error: continue from ghost context must target a ghost loop - } - } - - for i := 0 to 100 { - for j := 0 to u { - if i == t { - continue; // fine - } - } - } - - for i := 0 to 100 { - for j := 0 to u { - if i == t { - break continue; // error: continue from ghost context must target a ghost loop - } - } - } - - for i := 0 to 100 + u { - for j := 0 to u { - if i == t { - break continue; // fine - } - } - } - } - - method GhostContinueLabels(ghost t: int, ghost u: nat) - { - label Outer: - for i := 0 to 100 { - label Inner: - for j := 0 to u { - if j == t { - continue Inner; // fine - } else if j == 20 + t { - continue Outer; // error: continue from ghost context must target a ghost loop - } - } - } - } -} - -module LabelRegressions { - // The cases of if-case, while-case, and match statements are List's, which are essentially - // a BlockStmt but without the curly braces. Each Statement in such a List can have labels, so - // it's important to ResolveStatementWithLabels, not ResolveStatement. Alas, that was once not the - // case (pun intended). - // There's also something analogous going on in the Verifier, where lists of statements should call - // TrStmtList, not just call TrStmt on every Statement in the List. (See method LabelRegressions() - // in Test/comp/ForLoops-Compilation.dfy.) - method IfCaseRegression() { - if - case true => - label Loop: - for k := 0 to 10 { - continue Loop; - break Loop; - } - } - - method WhileCaseRegression() { - while - case true => - label Loop: - for k := 0 to 10 { - continue Loop; - break Loop; - } - } - - method Match() { - match (0, 0) - case (_, _) => - label Loop: - for k := 0 to 10 { - break Loop; - continue Loop; - } - } -} - -// --------------- regressions: using "this" in places where there is no enclosing class/type ------------------------------ - -module UseOfThis { - // The following uses of "this." once caused the resolver to crash. - - type {:badUseOfThis this.K} OpaqueType { // error: cannot use "this" here - const K: int - } - - newtype {:badUseOfThis this.x} Newtype = // error: cannot use "this" here - x: int | this.u // error: cannot use "this" here - witness this.u // error: cannot use "this" here - { - const K: int - } - - type {:badUseOfThis this.x} SynonymType = int // error: cannot use "this" here - - type {:badUseOfThis this.x} SubsetType = // error: cannot use "this" here - x: int | this.u // error: cannot use "this" here - witness this.u // error: cannot use "this" here - - trait {:badUseOfThis this.K} MyTrait { // error: cannot use "this" here - const K: int - } - - class {:badUseOfThis this.M} MyClass { // error: cannot use "this" here - const M: int - - var {:goodUseOfThis this.M} I: int - const {:goodUseOfThis this.M} J := 3 - method {:goodUseOfThis this.M} CM() - ensures {:goodUseOfThis this.M} true - ghost function {:goodUseOfThis this.M} CF(): int - ensures {:goodUseOfThis this.M} true - - static const {:badUseOfThis this.M} L := 3 // error: cannot use "this" here - static const N := this.M // error: cannot use "this" here - static method {:badUseOfThis this.M} SM() // error: cannot use "this" here - ensures {:badUseOfThis this.M} true // error: cannot use "this" here - static ghost function {:badUseOfThis this.M} SF(): int // error: cannot use "this" here - ensures {:badUseOfThis this.M} true // error: cannot use "this" here - } - - datatype Datatype = - | {:badUseOfThis this.K} DatatypeCtor // error: cannot use "this" here - { - const K: int - } -} - -module AutoInitTypeCheckRegression { - codatatype AutoStream = AutoNext(head: G, tail: AutoStream) - - function In(a: AutoStream): int // error: the argument to AutoStream is supposed to be an auto-init type - function Out(g: G): AutoStream // error: the argument to AutoStream is supposed to be an auto-init type - - method M(a: AutoStream) // error: the argument to AutoStream is supposed to be an auto-init type - method N(g: G) returns (a: AutoStream) // error: the argument to AutoStream is supposed to be an auto-init type -} - -module ClassConstructorTests { - class ClassWithConstructor { - var y: int - method NotTheOne() { } - constructor InitA() { } - constructor InitB() { y := 20; } - } - - class ClassWithoutConstructor { - method Init() modifies this { } - } - - method ConstructorTests() - { - var o := new object; // fine: does not have any constructors - - o := new ClassWithoutConstructor; // fine: don't need to call any particular method - o := new ClassWithoutConstructor.Init(); // this is also fine - - var c := new ClassWithConstructor.InitA(); - c := new ClassWithConstructor; // error: must call a constructor - c := new ClassWithConstructor.NotTheOne(); // error: must call a constructor, not an arbitrary method - c := new ClassWithConstructor.InitB(); - } - - class Y { - var data: int - constructor (x: int) - { - data := x; - } - method Test() { - var i := new Y(5); - i := new Y(7); - i := new Y; // error: the class has a constructor, so one must be used - var s := new Luci.Init(5); - s := new Luci.FromArray(null); - s := new Luci(false); - s := new Luci(true); - s := new Luci.M(); // error: there is a constructor, so one must be called - s := new Luci; // error: there is a constructor, so one must be called - var l := new Lamb; // fine, since there are no constructors (only methods) - l := new Lamb.Gwen(); - } - } - - class Luci { - constructor Init(y: int) { } - constructor (nameless: bool) { } - constructor FromArray(a: array) { } - method M() { } - } - - class Lamb { - method Jess() { } - method Gwen() { } - } -} - -module DefaultValues { - method Test() { - var r := 3 + MyFunction(7); - } - - function MyFunction(n: int, acc: int := n * false): int { // note, badly type default-value expression - 20 - } -} - -module IndexAdviceTypeRegressions { - type BV10 = x: bv10 | x != 999 witness 8 - - method Index(s: seq) - requires 2 < |s| - { - var k: bv10 := 2; - var a := s[k := "tjena"]; - - var K: BV10 := 2; - var A := s[K := "tjena"]; - - // rely on advice for the type - var i; - var b := s[i := "tjena"]; - } - - method Multi(matrix: array2) - requires matrix.Length0 == matrix.Length1 == 100 - { - var u := matrix[2, 5]; - var i := 2; - u := matrix[i, 5]; - var j := 2; - u := matrix[j, 5]; - j := j & j; - j := 16 as bv29; - // rely on advice for the type - var k; - u := matrix[k, 5]; - } - - method Array(arr: array) returns (r: real) - requires 10 <= arr.Length - { - var k: bv10 := 2; - r := arr[k]; - var K: BV10 := 2; - r := arr[K]; - } - - method Seq(s: seq) returns (r: real) - requires 10 <= |s| - { - var k: bv10 := 2; - r := s[k]; - var K: BV10 := 2; - r := s[K]; - } - - method Range(arr: array, s: seq) returns (r: seq) - requires 10 <= arr.Length <= |s| - { - var k: bv10 := 2; - var K: BV10 := 4; - r := arr[k..K]; - r := s[k..K]; - r := arr[k..]; - r := s[k..]; - r := arr[..K]; - r := s[..K]; - - // rely on advice for the type - var a, b, c, d; - r := arr[a..b]; - r := arr[c..]; - r := arr[..d]; - var w, x, y, z; - r := s[w..x]; - r := s[y..]; - r := s[..z]; - } - - method Bad(arr: array, s: seq, matrix: array2) returns (r: real, range: seq) { - r := arr[2.3]; // error: bad index type - r := s[2.3]; // error: bad index type - r := matrix[false, 2.3]; // error: bad index type - - range := arr[2.3..3.14]; // error (x2): bad index types - range := arr[2.3..]; // error: bad index type - range := arr[..3.14]; // error: bad index type - range := s[2.3..3.14]; // error (x2): bad index types - range := s[2.3..]; // error: bad index type - range := s[..3.14]; // error: bad index type - } -} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy.expect deleted file mode 100644 index e6086abf38e..00000000000 --- a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors.dfy.expect +++ /dev/null @@ -1,659 +0,0 @@ -ResolutionErrors.dfy(323,33): Warning: constructors no longer need 'this' to be listed in modifies clauses -ResolutionErrors.dfy(1879,15): Warning: constructors no longer need 'this' to be listed in modifies clauses -ResolutionErrors.dfy(13,18): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(25,13): Error: array selection requires an array2 (got array3) -ResolutionErrors.dfy(27,13): Error: array selection requires an array4 (got array) -ResolutionErrors.dfy(26,14): Error: element selection requires a sequence, array, multiset, or map (got array3) -ResolutionErrors.dfy(57,16): Error: accessing member 'X' requires an instance expression -ResolutionErrors.dfy(58,9): Error: unresolved identifier: F -ResolutionErrors.dfy(59,16): Error: accessing member 'F' requires an instance expression -ResolutionErrors.dfy(60,9): Error: unresolved identifier: G -ResolutionErrors.dfy(62,9): Error: unresolved identifier: M -ResolutionErrors.dfy(63,16): Error: accessing member 'M' requires an instance expression -ResolutionErrors.dfy(64,9): Error: unresolved identifier: N -ResolutionErrors.dfy(67,10): Error: non-function expression (of type int) is called with parameters -ResolutionErrors.dfy(68,16): Error: member 'z' does not exist in non-null type 'Global' -ResolutionErrors.dfy(51,15): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(88,16): Error: the name 'Benny' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.Benny') -ResolutionErrors.dfy(93,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') -ResolutionErrors.dfy(94,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') -ResolutionErrors.dfy(96,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') -ResolutionErrors.dfy(98,20): Error: wrong number of arguments (got 2, but datatype constructor 'David' expects 1: (x: int)) -ResolutionErrors.dfy(116,11): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(117,11): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(121,13): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(122,12): Error: actual out-parameter is required to be a ghost variable -ResolutionErrors.dfy(129,17): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(133,25): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(140,6): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(144,23): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(145,37): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(154,12): Error: only ghost methods can be called from this context -ResolutionErrors.dfy(244,29): Error: ghost-context break statement is not allowed to break out of non-ghost structure -ResolutionErrors.dfy(267,14): Error: ghost-context break statement is not allowed to break out of non-ghost loop -ResolutionErrors.dfy(281,14): Error: ghost-context break statement is not allowed to break out of non-ghost loop -ResolutionErrors.dfy(286,10): Error: return statement is not allowed in this context (because it is guarded by a specification-only expression) -ResolutionErrors.dfy(306,12): Error: label shadows an enclosing label -ResolutionErrors.dfy(311,10): Error: duplicate label -ResolutionErrors.dfy(313,10): Error: label shadows a dominating label -ResolutionErrors.dfy(329,11): Error: a constructor is allowed to be called only when an object is being allocated -ResolutionErrors.dfy(345,27): Error: arguments must have comparable types (got bool and int) -ResolutionErrors.dfy(343,18): Error: arguments must have comparable types (got int and DTD_List) -ResolutionErrors.dfy(344,18): Error: arguments must have comparable types (got DTD_List and int) -ResolutionErrors.dfy(358,17): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(382,7): Error: incorrect argument type at index 1 for method in-parameter 'b' (expected GenericClass, found GenericClass) (non-variant type parameter would require int = bool) -ResolutionErrors.dfy(396,13): Error: incorrect argument type at index 0 for datatype constructor parameter 'hd' (expected _T0, found int) -ResolutionErrors.dfy(397,9): Error: incorrect argument type at index 0 for datatype constructor parameter 'hd' (expected _T0, found int) -ResolutionErrors.dfy(406,8): Error: all lines in a calculation must have the same type (got int after bool) -ResolutionErrors.dfy(410,8): Error: all lines in a calculation must have the same type (got int after bool) -ResolutionErrors.dfy(413,8): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(413,8): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(414,12): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(414,12): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(415,12): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(415,12): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(413,8): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(413,8): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(419,12): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(419,12): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(418,8): Error: first argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(418,8): Error: second argument to ==> must be of type bool (instead got int) -ResolutionErrors.dfy(405,8): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) -ResolutionErrors.dfy(405,8): Error: arguments must have comparable types (got bool and int) -ResolutionErrors.dfy(405,8): Error: arguments must have comparable types (got bool and int) -ResolutionErrors.dfy(409,8): Error: arguments must have comparable types (got bool and int) -ResolutionErrors.dfy(409,8): Error: arguments must have comparable types (got bool and int) -ResolutionErrors.dfy(443,13): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(445,10): Error: a hint is not allowed to make heap updates -ResolutionErrors.dfy(447,10): Error: a hint is not allowed to update a variable it doesn't declare -ResolutionErrors.dfy(468,4): Error: More than one anonymous constructor -ResolutionErrors.dfy(481,15): Error: class Lamb does not have an anonymous constructor -ResolutionErrors.dfy(533,7): Error: RHS (of type List) not assignable to LHS (of type List) (covariant type parameter would require A <: B) -ResolutionErrors.dfy(538,7): Error: RHS (of type List) not assignable to LHS (of type List) (covariant type parameter would require A <: B) -ResolutionErrors.dfy(552,23): Error: type of case bodies do not agree (found Tree<_T1, _T0>, previous types Tree<_T0, _T1>) (covariant type parameter 0 would require _T1 <: _T0) -ResolutionErrors.dfy(564,30): Error: Wrong number of type arguments (0 instead of 2) passed to datatype: Tree -ResolutionErrors.dfy(579,20): Error: unresolved identifier: w -ResolutionErrors.dfy(598,8): Error: the type of this local variable is underspecified -ResolutionErrors.dfy(599,25): Error: the type of this variable is underspecified -ResolutionErrors.dfy(599,23): Error: type parameter 'T' (inferred to be '?') in the function call to 'P' could not be determined -ResolutionErrors.dfy(599,18): Error: type of bound variable 'z' could not be determined; please specify the type explicitly -ResolutionErrors.dfy(612,13): Error: a lemma is not allowed to use 'new' -ResolutionErrors.dfy(613,9): Error: a lemma is not allowed to use 'new' -ResolutionErrors.dfy(622,16): Error: only ghost methods can be called from this context -ResolutionErrors.dfy(629,15): Error: only ghost methods can be called from this context -ResolutionErrors.dfy(639,17): Error: a hint is not allowed to use 'new' -ResolutionErrors.dfy(654,14): Error: new allocation not supported in aggregate assignments -ResolutionErrors.dfy(661,14): Error: a forall statement is not allowed to use 'new' -ResolutionErrors.dfy(661,11): Error: assignment to array element is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(677,26): Error: second argument to "in" must be a set, multiset, or sequence with elements of type int, or a map with domain int (instead got bool) -ResolutionErrors.dfy(672,18): Error: second argument to "in" must be a set, multiset, or sequence with elements of type int, or a map with domain int (instead got ?) -ResolutionErrors.dfy(689,13): Error: lemmas are not allowed to have modifies clauses -ResolutionErrors.dfy(716,22): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(749,22): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(782,17): Error: the method returns 1 value but is assigned to 0 variable (all return values must be assigned) -ResolutionErrors.dfy(803,4): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(814,40): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(837,6): Error: RHS (of type B) not assignable to LHS (of type object) -ResolutionErrors.dfy(839,6): Error: RHS (of type B) not assignable to LHS (of type object) -ResolutionErrors.dfy(844,6): Error: RHS (of type G) not assignable to LHS (of type object) -ResolutionErrors.dfy(845,6): Error: RHS (of type Dt) not assignable to LHS (of type object) -ResolutionErrors.dfy(846,6): Error: RHS (of type CoDt) not assignable to LHS (of type object) -ResolutionErrors.dfy(838,6): Error: RHS (of type int) not assignable to LHS (of type object) -ResolutionErrors.dfy(855,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int) -ResolutionErrors.dfy(864,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(874,14): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(882,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(892,20): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(903,18): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(925,10): Error: LHS of assignment must denote a mutable field -ResolutionErrors.dfy(926,10): Error: LHS of assignment must denote a mutable field -ResolutionErrors.dfy(927,9): Error: cannot assign to a range of array elements (try the 'forall' statement) -ResolutionErrors.dfy(928,9): Error: cannot assign to a range of array elements (try the 'forall' statement) -ResolutionErrors.dfy(929,5): Error: cannot assign to a range of array elements (try the 'forall' statement) -ResolutionErrors.dfy(930,5): Error: cannot assign to a range of array elements (try the 'forall' statement) -ResolutionErrors.dfy(919,4): Error: LHS of array assignment must denote an array element (found seq) -ResolutionErrors.dfy(920,4): Error: LHS of array assignment must denote an array element (found seq) -ResolutionErrors.dfy(942,11): Error: unresolved identifier: s -ResolutionErrors.dfy(961,18): Error: member '3' does not exist in datatype '_tuple#3' -ResolutionErrors.dfy(961,28): Error: member 'x' does not exist in datatype '_tuple#2' -ResolutionErrors.dfy(960,18): Error: condition is expected to be of type bool, but is int -ResolutionErrors.dfy(954,39): Error: RHS (of type (int, real, int)) not assignable to LHS (of type (int, real, int, real)) -ResolutionErrors.dfy(953,34): Error: RHS (of type (int, int, real)) not assignable to LHS (of type (int, real, int)) (covariant type parameter 1 would require int <: real) -ResolutionErrors.dfy(953,34): Error: RHS (of type (int, int, real)) not assignable to LHS (of type (int, real, int)) (covariant type parameter 2 would require real <: int) -ResolutionErrors.dfy(985,12): Error: type of left argument to % (int) must agree with the result type (real) -ResolutionErrors.dfy(985,12): Error: arguments to % must be integer-numeric or bitvector types (got real) -ResolutionErrors.dfy(984,19): Error: type of right argument to / (int) must agree with the result type (real) -ResolutionErrors.dfy(1012,11): Error: Wrong number of type arguments (2 instead of 1) passed to non-null type: array3 -ResolutionErrors.dfy(1013,11): Error: Wrong number of type arguments (2 instead of 1) passed to non-null type: C -ResolutionErrors.dfy(1024,7): Error: duplicate name of top-level declaration: BadSyn2 -ResolutionErrors.dfy(1023,7): Related location -ResolutionErrors.dfy(1021,17): Error: Wrong number of type arguments (0 instead of 1) passed to datatype: List -ResolutionErrors.dfy(1022,17): Error: Type or type parameter is not declared in this scope: badName (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1023,22): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1030,7): Error: type-synonym cycle: A -> A -ResolutionErrors.dfy(1033,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A -ResolutionErrors.dfy(1037,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A -ResolutionErrors.dfy(1046,11): Warning: because of cyclic dependencies among constructor argument types, no instances of datatype 'D' can be constructed -ResolutionErrors.dfy(1043,8): Error: class 'C' with fields without known initializers, like 'a' of type 'A', must declare a constructor -ResolutionErrors.dfy(1049,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A -ResolutionErrors.dfy(1054,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A -ResolutionErrors.dfy(1059,7): Error: type-synonym cycle: A -> A -ResolutionErrors.dfy(1066,25): Error: unresolved identifier: x -ResolutionErrors.dfy(1069,22): Error: unresolved identifier: x -ResolutionErrors.dfy(1072,25): Error: unresolved identifier: x -ResolutionErrors.dfy(1074,21): Error: unresolved identifier: x -ResolutionErrors.dfy(1076,21): Error: unresolved identifier: x -ResolutionErrors.dfy(1079,21): Error: unresolved identifier: x -ResolutionErrors.dfy(1086,35): Error: Wrong number of type arguments (2 instead of 1) passed to abstract type: P -ResolutionErrors.dfy(1098,13): Error: Type or type parameter is not declared in this scope: BX (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1108,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require bool = int) -ResolutionErrors.dfy(1113,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require B = A) -ResolutionErrors.dfy(1118,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require int = A) -ResolutionErrors.dfy(1119,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require A = int) -ResolutionErrors.dfy(1124,13): Error: arguments must have comparable types (got P and P) -ResolutionErrors.dfy(1125,13): Error: arguments must have comparable types (got P and P) -ResolutionErrors.dfy(1126,13): Error: arguments must have comparable types (got P and P) -ResolutionErrors.dfy(1137,13): Error: new can be applied only to class types (got JJ) -ResolutionErrors.dfy(1145,37): Error: a set comprehension involved in a function definition is not allowed to depend on the set of allocated references, but values of 'o' (of type 'object') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(1146,31): Error: a set comprehension involved in a function definition is not allowed to depend on the set of allocated references, but values of 'o' (of type 'object') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(1152,27): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'o' -ResolutionErrors.dfy(1160,15): Error: arguments to / must be numeric or bitvector types (got set) -ResolutionErrors.dfy(1167,20): Error: a call to a possibly non-terminating method is allowed only if the calling method is also declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(1182,16): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(1247,13): Error: type parameter 'PT' (inferred to be '?') in the function call to 'P' could not be determined -ResolutionErrors.dfy(1248,14): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1248,19): Error: type parameter 'QT' (inferred to be '?') in the function call to 'Q' could not be determined -ResolutionErrors.dfy(1248,20): Error: the type of this expression is underspecified -ResolutionErrors.dfy(1249,4): Error: type parameter 'MT' (inferred to be '?') to the method 'M' could not be determined -ResolutionErrors.dfy(1250,8): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1250,13): Error: type parameter 'NT' (inferred to be '?') to the method 'N' could not be determined -ResolutionErrors.dfy(1251,8): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1252,8): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1253,8): Error: the type of this local variable is underspecified -ResolutionErrors.dfy(1254,8): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1255,8): Error: the type of this local variable is underspecified -ResolutionErrors.dfy(1259,26): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1259,21): Error: type of bound variable 's' could not be determined; please specify the type explicitly -ResolutionErrors.dfy(1260,31): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1260,21): Error: type of bound variable 's' could not be determined; please specify the type explicitly -ResolutionErrors.dfy(1261,30): Error: the type of this variable is underspecified -ResolutionErrors.dfy(1261,21): Error: type of bound variable 'c' could not be determined; please specify the type explicitly -ResolutionErrors.dfy(1268,8): Error: the type of this local variable is underspecified -ResolutionErrors.dfy(1271,29): Error: type of bound variable 'c' could not be determined; please specify the type explicitly -ResolutionErrors.dfy(1282,21): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1283,24): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1320,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (y) -ResolutionErrors.dfy(1337,18): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(1365,15): Error: The name Inner ambiguously refers to a type in one of the modules A, B (try qualifying the type name with the module name) -ResolutionErrors.dfy(1375,29): Error: ghost variables such as tmp are allowed only in specification contexts. tmp was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(1377,49): Error: ghost variables such as a0 are allowed only in specification contexts. a0 was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(1377,54): Error: ghost variables such as a1 are allowed only in specification contexts. a1 was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(1386,33): Error: set argument type must support equality (got (int, int -> bool)) -ResolutionErrors.dfy(1398,11): Error: name of type (X) is used as a variable -ResolutionErrors.dfy(1398,16): Error: name of type (X) is used as a variable -ResolutionErrors.dfy(1399,11): Error: name of module (Y) is used as a variable -ResolutionErrors.dfy(1399,16): Error: name of module (Y) is used as a variable -ResolutionErrors.dfy(1400,11): Error: name of type (X) is used as a variable -ResolutionErrors.dfy(1401,11): Error: name of module (Y) is used as a variable -ResolutionErrors.dfy(1406,16): Error: name of type (X) is used as a variable -ResolutionErrors.dfy(1407,16): Error: name of module (Y) is used as a variable -ResolutionErrors.dfy(1408,4): Error: name of type (X) is used as a variable -ResolutionErrors.dfy(1409,4): Error: name of module (Y) is used as a variable -ResolutionErrors.dfy(1419,9): Error: type of RHS of assign-such-that statement must be boolean (got int) -ResolutionErrors.dfy(1418,11): Error: type of left argument to + (int) must agree with the result type (bool) -ResolutionErrors.dfy(1418,11): Error: type of right argument to + (int) must agree with the result type (bool) -ResolutionErrors.dfy(1418,11): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) -ResolutionErrors.dfy(1420,13): Error: type of RHS of assign-such-that statement must be boolean (got int) -ResolutionErrors.dfy(1423,15): Error: type of left argument to + (int) must agree with the result type (bool) -ResolutionErrors.dfy(1423,15): Error: type of right argument to + (int) must agree with the result type (bool) -ResolutionErrors.dfy(1423,15): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) -ResolutionErrors.dfy(1439,29): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1441,17): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1466,20): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1475,24): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1488,18): Error: assignment to non-ghost field is not allowed in this context, because this is a ghost function -ResolutionErrors.dfy(1489,10): Error: a hint is not allowed to make heap updates -ResolutionErrors.dfy(1490,20): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1493,21): Error: a loop in a hint is not allowed to use 'modifies' clauses -ResolutionErrors.dfy(1510,24): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1517,18): Error: assignment to non-ghost field is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(1518,10): Error: a hint is not allowed to make heap updates -ResolutionErrors.dfy(1519,11): Error: in a hint, calls are allowed only to lemmas -ResolutionErrors.dfy(1522,21): Error: a loop in a hint is not allowed to use 'modifies' clauses -ResolutionErrors.dfy(1545,29): Error: in a statement expression, calls are allowed only to lemmas -ResolutionErrors.dfy(1557,16): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating -ResolutionErrors.dfy(1575,12): Error: a 'break break break' statement is allowed only in contexts with 3 enclosing loops, but the current context only has 2 -ResolutionErrors.dfy(1587,20): Error: a ghost field is allowed only in specification contexts -ResolutionErrors.dfy(1594,9): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(1600,4): Error: non-ghost variable cannot be assigned a value that depends on a ghost -ResolutionErrors.dfy(1617,8): Error: print statement is not allowed in this context (because this is a ghost method or because the statement is guarded by a specification-only expression) -ResolutionErrors.dfy(1626,29): Error: ghost variables such as z are allowed only in specification contexts. z was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(1634,10): Error: the type of the bound variable 't' could not be determined -ResolutionErrors.dfy(1634,10): Error: the type of the bound variable 't' could not be determined -ResolutionErrors.dfy(1634,10): Error: the type of the bound variable 't' could not be determined -ResolutionErrors.dfy(1634,10): Error: the type of the bound variable 't' could not be determined -ResolutionErrors.dfy(1652,15): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(1654,10): Error: assignment to non-ghost field is not allowed in this context, because this is a ghost method -ResolutionErrors.dfy(1679,15): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) -ResolutionErrors.dfy(1681,25): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(1682,35): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(1692,4): Error: 'decreases *' is not allowed on ghost loops -ResolutionErrors.dfy(1696,8): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(1706,4): Error: 'decreases *' is not allowed on ghost loops -ResolutionErrors.dfy(1710,29): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression -ResolutionErrors.dfy(1718,21): Error: the type of the bound variable 'u' could not be determined -ResolutionErrors.dfy(1719,23): Error: the type of the bound variable 'u' could not be determined -ResolutionErrors.dfy(1722,27): Error: the type of the bound variable 'u' could not be determined -ResolutionErrors.dfy(1726,40): Error: the type of the bound variable 'u' could not be determined -ResolutionErrors.dfy(1728,38): Error: the type of the bound variable 'u' could not be determined -ResolutionErrors.dfy(1772,19): Error: Type or type parameter is not declared in this scope: realint (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) -ResolutionErrors.dfy(1780,44): Error: range of comprehension must be of type bool (instead got C) -ResolutionErrors.dfy(1787,21): Error: Duplicate member name: P -ResolutionErrors.dfy(1789,17): Error: Duplicate member name: L -ResolutionErrors.dfy(1797,4): Error: second argument to && must be of type bool (instead got int) -ResolutionErrors.dfy(1795,17): Error: Function body type mismatch (expected int, got bool) -ResolutionErrors.dfy(1801,4): Error: second argument to && must be of type bool (instead got int) -ResolutionErrors.dfy(1805,4): Error: second argument to || must be of type bool (instead got int) -ResolutionErrors.dfy(1803,17): Error: Function body type mismatch (expected int, got bool) -ResolutionErrors.dfy(1809,4): Error: second argument to || must be of type bool (instead got int) -ResolutionErrors.dfy(1813,8): Error: second argument to || must be of type bool (instead got int) -ResolutionErrors.dfy(1811,17): Error: Function body type mismatch (expected int, got bool) -ResolutionErrors.dfy(1834,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1835,11): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1835,15): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1836,16): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1837,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1840,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1841,19): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1842,8): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1844,8): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1845,13): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1846,13): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1847,14): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields -ResolutionErrors.dfy(1878,15): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(1880,16): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(1896,8): Error: return statement is not allowed before 'new;' in a constructor -ResolutionErrors.dfy(1922,5): Error: type parameter (G) passed to method P must support auto-initialization (got F) (perhaps try declaring type parameter 'F' on line 1915 as 'F(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(1924,5): Error: type parameter (G) passed to method P must support auto-initialization (got Y) (perhaps try declaring abstract type 'Y' on line 1912 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(1961,5): Error: type parameter (G) passed to method P must support auto-initialization (got SixOrMore) -ResolutionErrors.dfy(1962,5): Error: type parameter (G) passed to method P must support auto-initialization (got AnotherSixOrMore) -ResolutionErrors.dfy(1963,5): Error: type parameter (G) passed to method P must support auto-initialization (got MySixOrMore) -ResolutionErrors.dfy(1964,5): Error: type parameter (G) passed to method P must support auto-initialization (got UnclearA) -ResolutionErrors.dfy(1980,17): Error: type parameter (G) passed to method P must support auto-initialization (got SSixOrMore) -ResolutionErrors.dfy(1981,24): Error: type parameter (G) passed to method P must support auto-initialization (got SAnotherSixOrMore) -ResolutionErrors.dfy(1982,19): Error: type parameter (G) passed to method P must support auto-initialization (got SMySixOrMore) -ResolutionErrors.dfy(1983,16): Error: type parameter (G) passed to method P must support auto-initialization (got SUnclearA) -ResolutionErrors.dfy(2013,7): Error: type 'B2', which does not support auto-initialization, is used to refine an abstract type that expects auto-initialization -ResolutionErrors.dfy(2014,7): Error: type 'B3', which does not support auto-initialization, is used to refine an abstract type that expects auto-initialization -ResolutionErrors.dfy(2018,7): Error: type 'C3', which may be empty, is used to refine an abstract type expected to be nonempty -ResolutionErrors.dfy(2048,4): Error: type parameter 'Q' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2051,4): Error: type parameter 'R' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2038,7): Error: type declaration 'A1' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2039,7): Error: type declaration 'A2' is not allowed to change the requirement of being nonempty -ResolutionErrors.dfy(2040,7): Error: type declaration 'B0' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2042,7): Error: type declaration 'B2' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2043,7): Error: type declaration 'C0' is not allowed to change the requirement of being nonempty -ResolutionErrors.dfy(2044,7): Error: type declaration 'C1' is not allowed to change the requirement of supporting auto-initialization -ResolutionErrors.dfy(2034,5): Error: type parameter (G) passed to method P must support auto-initialization (got Z.XYZ) -ResolutionErrors.dfy(2097,33): Error: type of yield-parameter must support auto-initialization (got 'T') -ResolutionErrors.dfy(2105,13): Error: a ghost function is allowed only in specification contexts -ResolutionErrors.dfy(2119,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2120,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2122,19): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F' is not allowed to use ORDINAL -ResolutionErrors.dfy(2123,9): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2124,9): Error: type parameter 'G' (passed in as '(char, ORDINAL)') to function call 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2125,18): Error: type parameter 'G' (passed in as 'ORDINAL') to the function 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2126,4): Error: type parameter 'G' (passed in as 'ORDINAL') to the method 'ParameterizedMethod' is not allowed to use ORDINAL -ResolutionErrors.dfy(2130,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2132,14): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2133,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2134,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2135,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2137,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2140,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2140,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2141,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2141,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2141,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2141,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2143,25): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2169,36): Error: Name of yield-parameter is used by another member of the iterator: u -ResolutionErrors.dfy(2170,37): Error: Name of implicit yield-history variable 'u' is already used by another member of the iterator -ResolutionErrors.dfy(2194,15): Error: To access members of class 'MyClass', write 'MyClass', not 'MyClass?' -ResolutionErrors.dfy(2205,15): Error: name of type (MyClass) is used as a variable -ResolutionErrors.dfy(2206,15): Error: name of type (MyClass?) is used as a variable -ResolutionErrors.dfy(2219,17): Error: To access members of class 'MyClass', write 'MyClass', not 'MyClass?' -ResolutionErrors.dfy(2230,17): Error: name of type (MyClass) is used as a variable -ResolutionErrors.dfy(2231,17): Error: name of type (MyClass?) is used as a variable -ResolutionErrors.dfy(2240,8): Error: static non-ghost const field 'X' of type 'Odd' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2243,8): Error: static non-ghost const field 'Y' of type 'MyClass' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2244,14): Error: static ghost const field 'Y'' of type 'MyClass' (which may be empty) must give a defining value -ResolutionErrors.dfy(2248,17): Error: static non-ghost const field 'W' of type 'MyClass' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2250,17): Error: static non-ghost const field 'O' of type 'Odd' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2256,17): Error: static non-ghost const field 'W' of type 'MyClass' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2258,17): Error: static non-ghost const field 'O' of type 'Odd' (which does not have a default compiled value) must give a defining value -ResolutionErrors.dfy(2266,8): Error: class 'Instance' with fields without known initializers, like 'w' of type 'MyClass', must declare a constructor -ResolutionErrors.dfy(2272,8): Error: class 'GhostCl' with fields without known initializers, like 'z' of type 'MyClass', must declare a constructor -ResolutionErrors.dfy(2282,6): Error: LHS of assignment must denote a mutable field -ResolutionErrors.dfy(2283,10): Error: LHS of assignment must denote a mutable field -ResolutionErrors.dfy(2290,6): Error: LHS of assignment must denote a mutable field -ResolutionErrors.dfy(2297,8): Error: LHS of assignment must denote a mutable field of 'this' -ResolutionErrors.dfy(2314,18): Error: a ghost const field is allowed only in specification contexts -ResolutionErrors.dfy(2311,20): Error: a ghost const field is allowed only in specification contexts -ResolutionErrors.dfy(2330,9): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' -ResolutionErrors.dfy(2333,9): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' -ResolutionErrors.dfy(2341,22): Error: a set comprehension involved in a const field definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(2339,24): Error: a set comprehension involved in a const field definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(2342,36): Error: a set comprehension involved in a newtype definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(2343,34): Error: a set comprehension involved in a subset type definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(2348,18): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'c' -ResolutionErrors.dfy(2349,18): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'c' -ResolutionErrors.dfy(2355,19): Error: a set comprehension involved in a predicate definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) -ResolutionErrors.dfy(2363,35): Error: a const field definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2361,37): Error: a const field definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2364,49): Error: a newtype definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2365,47): Error: a subset type definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2370,32): Error: a predicate definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2378,39): Error: a const field definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2376,41): Error: a const field definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2379,53): Error: a newtype definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2380,51): Error: a subset type definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2385,36): Error: a predicate definition is not allowed to depend on the set of allocated references -ResolutionErrors.dfy(2392,31): Error: the value returned by an abstemious function must come from invoking a co-constructor -ResolutionErrors.dfy(2428,31): Error: the value returned by an abstemious function must come from invoking a co-constructor -ResolutionErrors.dfy(2435,32): Error: an abstemious function is allowed to invoke a codatatype destructor only on its parameters -ResolutionErrors.dfy(2440,12): Error: an abstemious function is allowed to codatatype-match only on its parameters -ResolutionErrors.dfy(2447,9): Error: an abstemious function is not allowed to check codatatype equality -ResolutionErrors.dfy(2449,14): Error: an abstemious function is not allowed to check codatatype equality -ResolutionErrors.dfy(2474,19): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F' is not allowed to use ORDINAL -ResolutionErrors.dfy(2475,13): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2476,13): Error: type parameter 'G' (passed in as '(char, ORDINAL)') to function call 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2477,18): Error: type parameter 'G' (passed in as 'ORDINAL') to the function 'F'' is not allowed to use ORDINAL -ResolutionErrors.dfy(2478,4): Error: type parameter 'G' (passed in as 'ORDINAL') to the lemma 'ParameterizedLemma' is not allowed to use ORDINAL -ResolutionErrors.dfy(2479,18): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2480,18): Error: type of bound variable 'r' ('(ORDINAL, int)') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2481,26): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2482,17): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2483,18): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2484,24): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2495,25): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2496,17): Error: type of bound variable 'yt' ('ORDINAL') is not allowed to use type ORDINAL -ResolutionErrors.dfy(2493,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2493,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2492,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2489,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2487,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2486,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2485,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2484,14): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2482,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2472,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2471,8): Error: an ORDINAL type is not allowed to be used as a type argument -ResolutionErrors.dfy(2528,11): Error: no label 'L' in scope at this time -ResolutionErrors.dfy(2531,10): Error: label shadows a dominating label -ResolutionErrors.dfy(2539,10): Error: label shadows a dominating label -ResolutionErrors.dfy(2611,23): Error: no label 'Treasure' in scope at this time -ResolutionErrors.dfy(2612,19): Error: no label 'WonderfulLabel' in scope at this time -ResolutionErrors.dfy(2614,23): Error: no label 'FutureLabel' in scope at this time -ResolutionErrors.dfy(2623,13): Error: no label 'Treasure' in scope at this time -ResolutionErrors.dfy(2624,13): Error: no label 'WonderfulLabel' in scope at this time -ResolutionErrors.dfy(2626,13): Error: no label 'FutureLabel' in scope at this time -ResolutionErrors.dfy(2635,13): Error: no label 'Treasure' in scope at this time -ResolutionErrors.dfy(2636,13): Error: no label 'WonderfulLabel' in scope at this time -ResolutionErrors.dfy(2638,13): Error: no label 'FutureLabel' in scope at this time -ResolutionErrors.dfy(2667,26): Error: type parameter (G) passed to method GimmieOne must support auto-initialization (got Yt) -ResolutionErrors.dfy(2680,50): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(2706,11): Error: name of type (Cache) is used as a variable -ResolutionErrors.dfy(2706,17): Error: incorrect type for selection into ? (got X) -ResolutionErrors.dfy(2714,13): Warning: the quantifier has the form 'exists x :: A ==> B', which most often is a typo for 'exists x :: A && B'; if you think otherwise, rewrite as 'exists x :: (A ==> B)' or 'exists x :: !A || B' to suppress this warning -ResolutionErrors.dfy(2724,13): Warning: the quantifier has the form 'exists x :: A ==> B', which most often is a typo for 'exists x :: A && B'; if you think otherwise, rewrite as 'exists x :: (A ==> B)' or 'exists x :: !A || B' to suppress this warning -ResolutionErrors.dfy(2757,13): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2758,13): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2760,11): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2781,4): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2791,15): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2792,22): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2798,15): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2799,22): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(2816,23): Error: type of left argument to * (int) must agree with the result type (bool) -ResolutionErrors.dfy(2816,23): Error: type of * must be of a numeric type, bitvector type, or a set-like type (instead got bool) -ResolutionErrors.dfy(2815,13): Error: not resolving module 'V' because there were errors in resolving its nested module 'W' -ResolutionErrors.dfy(2825,21): Error: type of left argument to * (int) must agree with the result type (bool) -ResolutionErrors.dfy(2825,21): Error: type of * must be of a numeric type, bitvector type, or a set-like type (instead got bool) -ResolutionErrors.dfy(2824,11): Error: not resolving module 'V' because there were errors in resolving its nested module 'W' -ResolutionErrors.dfy(2811,9): Error: not resolving module 'ErrorsFromNestedModules' because there were errors in resolving its nested module 'U' -ResolutionErrors.dfy(2833,11): Error: duplicate name of top-level declaration: G -ResolutionErrors.dfy(2836,10): Related location -ResolutionErrors.dfy(2839,11): Error: duplicate name of top-level declaration: H -ResolutionErrors.dfy(2837,10): Related location -ResolutionErrors.dfy(2855,4): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2857,20): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2866,6): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2867,4): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2876,6): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2877,8): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2878,4): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2890,13): Error: expression is not allowed to refer to constant field x -ResolutionErrors.dfy(2892,22): Error: expression is not allowed to refer to constant field x -ResolutionErrors.dfy(2903,13): Error: arguments must be of a set or multiset type (got map) -ResolutionErrors.dfy(2913,4): Error: expect statement is not allowed in this context (because this is a ghost method or because the statement is guarded by a specification-only expression) -ResolutionErrors.dfy(2919,11): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' -ResolutionErrors.dfy(2920,21): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(2931,15): Error: RHS (of type X) not assignable to LHS (of type X) -ResolutionErrors.dfy(2935,10): Error: type of corresponding source/RHS (X) does not match type of bound variable (X) -ResolutionErrors.dfy(2963,16): Error: sequence update requires the value to have the element type of the sequence (got Trait) -ResolutionErrors.dfy(2967,11): Error: multiset update requires domain element to be of type Elem (got Trait) -ResolutionErrors.dfy(2972,13): Error: map update requires domain element to be of type Elem (got Trait) -ResolutionErrors.dfy(2974,18): Error: map update requires the value to have the range type Elem (got Trait) -ResolutionErrors.dfy(2999,5): Error: type parameter (F) passed to method Q must support auto-initialization (got Y) (perhaps try declaring type parameter 'Y' on line 2994 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3002,5): Error: type parameter (F) passed to method Q must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 2994 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3003,5): Error: type parameter (G) passed to method P must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 2994 as 'Z(00)', which says it can only be instantiated with a nonempty type) -ResolutionErrors.dfy(3012,9): Error: type parameter (F) passed to function FQ must support auto-initialization (got Y) (perhaps try declaring type parameter 'Y' on line 3007 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3015,9): Error: type parameter (F) passed to function FQ must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 3007 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3016,9): Error: type parameter (G) passed to function FP must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 3007 as 'Z(00)', which says it can only be instantiated with a nonempty type) -ResolutionErrors.dfy(3033,9): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) -ResolutionErrors.dfy(3034,9): Error: a call to a ghost predicate is allowed only in specification contexts (consider declaring the predicate without the 'ghost' keyword) -ResolutionErrors.dfy(3035,9): Error: a call to a least predicate is allowed only in specification contexts -ResolutionErrors.dfy(3036,9): Error: a call to a greatest predicate is allowed only in specification contexts -ResolutionErrors.dfy(3037,9): Error: a call to a twostate function is allowed only in specification contexts -ResolutionErrors.dfy(3038,9): Error: a call to a twostate predicate is allowed only in specification contexts -ResolutionErrors.dfy(3055,9): Error: function 'F0' expects 0 type arguments (got 1) -ResolutionErrors.dfy(3057,9): Error: function 'F2' expects 2 type arguments (got 1) -ResolutionErrors.dfy(3058,9): Error: function 'F0' expects 0 type arguments (got 2) -ResolutionErrors.dfy(3059,9): Error: function 'F1' expects 1 type argument (got 2) -ResolutionErrors.dfy(3067,4): Error: method 'M0' expects 0 type arguments (got 1) -ResolutionErrors.dfy(3069,4): Error: method 'M2' expects 2 type arguments (got 1) -ResolutionErrors.dfy(3070,4): Error: method 'M0' expects 0 type arguments (got 2) -ResolutionErrors.dfy(3071,4): Error: method 'M1' expects 1 type argument (got 2) -ResolutionErrors.dfy(3084,10): Error: ghost variables such as c are allowed only in specification contexts. c was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3088,14): Error: ghost variables such as t are allowed only in specification contexts. t was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3093,10): Error: ghost variables such as a' are allowed only in specification contexts. a' was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3094,10): Error: ghost variables such as c' are allowed only in specification contexts. c' was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3118,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3123,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 3114 as 'Z(00)', which says it can only be instantiated with a nonempty type) -ResolutionErrors.dfy(3125,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3126,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) -ResolutionErrors.dfy(3130,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 3114 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3134,9): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3137,9): Error: type parameter (T) passed to function MustSupportEquality must support equality (got Z) (perhaps try declaring type parameter 'Z' on line 3114 as 'Z(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3142,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3144,9): Error: type parameter (T) passed to function NoReferences must support no references (got Z) (perhaps try declaring type parameter 'Z' on line 3114 as 'Z(!new)', which says it can only be instantiated with a type that contains no references) -ResolutionErrors.dfy(3151,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3157,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3172,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3181,12): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3179 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3183,7): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3179 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3186,12): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) -ResolutionErrors.dfy(3188,7): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) -ResolutionErrors.dfy(3191,19): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) -ResolutionErrors.dfy(3199,13): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3198 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3204,13): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3198 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3206,23): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3198 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3208,14): Error: type parameter 0 (T) passed to type QuadEq must support equality (got seq) -ResolutionErrors.dfy(3209,15): Error: type parameter 1 (U) passed to type QuadEq must support equality (got seq) -ResolutionErrors.dfy(3218,15): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3215 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3227,15): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 3215 as 'T(==)', which says it can only be instantiated with a type that supports equality) -ResolutionErrors.dfy(3245,11): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3246,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3247,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) -ResolutionErrors.dfy(3248,11): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3249,11): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3253,47): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3254,12): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3258,11): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3259,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3260,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) -ResolutionErrors.dfy(3261,11): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3262,11): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3266,50): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3267,12): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3270,12): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3274,12): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3279,53): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3282,63): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3299,6): Error: ghost variables such as m are allowed only in specification contexts. m was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3303,8): Error: non-ghost variable cannot be assigned a value that depends on a ghost -ResolutionErrors.dfy(3307,18): Error: ghost variables such as m are allowed only in specification contexts. m was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3333,6): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3339,6): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3340,6): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) -ResolutionErrors.dfy(3347,6): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) -ResolutionErrors.dfy(3354,6): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3360,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) -ResolutionErrors.dfy(3366,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) -ResolutionErrors.dfy(3381,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) -ResolutionErrors.dfy(3429,7): Error: ghost variables such as p are allowed only in specification contexts. p was inferred to be ghost based on its declaration or initialization. -ResolutionErrors.dfy(3443,21): Error: a loop in a hint is not allowed to use 'modifies' clauses -ResolutionErrors.dfy(3450,21): Error: a loop in a hint is not allowed to use 'modifies' clauses -ResolutionErrors.dfy(3479,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got multiset) -ResolutionErrors.dfy(3480,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got map) -ResolutionErrors.dfy(3481,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got map) -ResolutionErrors.dfy(3482,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got imap) -ResolutionErrors.dfy(3487,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int) -ResolutionErrors.dfy(3488,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got set) -ResolutionErrors.dfy(3489,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got iset) -ResolutionErrors.dfy(3490,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got seq) -ResolutionErrors.dfy(3495,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> int) -ResolutionErrors.dfy(3496,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> object) -ResolutionErrors.dfy(3497,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> set) -ResolutionErrors.dfy(3498,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> seq) -ResolutionErrors.dfy(3499,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got set -> int) -ResolutionErrors.dfy(3518,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3519,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3520,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got imap) -ResolutionErrors.dfy(3525,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int) -ResolutionErrors.dfy(3526,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got set) -ResolutionErrors.dfy(3527,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got iset) -ResolutionErrors.dfy(3528,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got seq) -ResolutionErrors.dfy(3537,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> int) -ResolutionErrors.dfy(3538,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> object) -ResolutionErrors.dfy(3539,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> set) -ResolutionErrors.dfy(3540,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> iset) -ResolutionErrors.dfy(3541,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> seq) -ResolutionErrors.dfy(3542,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got set -> int) -ResolutionErrors.dfy(3543,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> multiset) -ResolutionErrors.dfy(3544,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> map) -ResolutionErrors.dfy(3559,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3560,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3561,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got imap) -ResolutionErrors.dfy(3566,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int) -ResolutionErrors.dfy(3567,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got set) -ResolutionErrors.dfy(3568,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got iset) -ResolutionErrors.dfy(3569,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got seq) -ResolutionErrors.dfy(3577,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> int) -ResolutionErrors.dfy(3578,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> object) -ResolutionErrors.dfy(3579,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> set) -ResolutionErrors.dfy(3580,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> iset) -ResolutionErrors.dfy(3581,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> seq) -ResolutionErrors.dfy(3582,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got set -> int) -ResolutionErrors.dfy(3583,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> multiset) -ResolutionErrors.dfy(3584,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> map) -ResolutionErrors.dfy(3600,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3601,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got map) -ResolutionErrors.dfy(3602,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got imap) -ResolutionErrors.dfy(3608,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int) -ResolutionErrors.dfy(3609,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got set) -ResolutionErrors.dfy(3610,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got iset) -ResolutionErrors.dfy(3611,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got seq) -ResolutionErrors.dfy(3620,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int -> int) -ResolutionErrors.dfy(3621,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int -> object) -ResolutionErrors.dfy(3625,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got set -> int) -ResolutionErrors.dfy(3627,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got bool -> map) -ResolutionErrors.dfy(3646,12): Error: a 'break break break continue' statement is allowed only in contexts with 4 enclosing loops, but the current context only has 3 -ResolutionErrors.dfy(3683,27): Error: continue label must designate a loop: X -ResolutionErrors.dfy(3685,27): Error: continue label must designate a loop: Y0 -ResolutionErrors.dfy(3687,27): Error: continue label must designate a loop: Y1 -ResolutionErrors.dfy(3689,27): Error: continue label must designate a loop: Z -ResolutionErrors.dfy(3699,4): Error: a non-labeled 'break' statement is allowed only in loops -ResolutionErrors.dfy(3703,4): Error: a non-labeled 'continue' statement is allowed only in loops -ResolutionErrors.dfy(3709,19): Error: a non-labeled 'break' statement is allowed only in loops -ResolutionErrors.dfy(3710,19): Error: a non-labeled 'continue' statement is allowed only in loops -ResolutionErrors.dfy(3712,28): Error: continue label must designate a loop: X -ResolutionErrors.dfy(3729,21): Error: a non-labeled 'break' statement is allowed only in loops -ResolutionErrors.dfy(3730,21): Error: a non-labeled 'continue' statement is allowed only in loops -ResolutionErrors.dfy(3731,27): Error: break label is undefined or not in scope: L -ResolutionErrors.dfy(3732,30): Error: continue label is undefined or not in scope: L -ResolutionErrors.dfy(3745,8): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop -ResolutionErrors.dfy(3754,8): Error: ghost-context break statement is not allowed to break out of non-ghost loop -ResolutionErrors.dfy(3762,8): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop -ResolutionErrors.dfy(3777,10): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop -ResolutionErrors.dfy(3800,10): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop -ResolutionErrors.dfy(3856,13): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3856,18): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3865,13): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3865,18): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3883,22): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3851,22): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3851,27): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3857,12): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3857,17): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3855,25): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3855,30): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3862,22): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3862,27): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3866,12): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3866,17): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3864,22): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3864,27): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3868,23): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3868,28): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3882,32): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3885,29): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3884,33): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3887,29): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3886,41): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3872,23): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3872,28): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3891,21): Error: 'this' is not allowed in a 'static' context -ResolutionErrors.dfy(3891,26): Error: type of the receiver is not fully determined at this program point -ResolutionErrors.dfy(3900,20): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 3900 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3901,25): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 3901 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3903,17): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 3903 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3904,32): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 3904 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) -ResolutionErrors.dfy(3927,6): Error: when allocating an object of type 'ClassWithConstructor', one of its constructor methods must be called -ResolutionErrors.dfy(3928,6): Error: when allocating an object of type 'ClassWithConstructor', one of its constructor methods must be called -ResolutionErrors.dfy(3941,8): Error: when allocating an object of type 'Y', one of its constructor methods must be called -ResolutionErrors.dfy(3946,8): Error: when allocating an object of type 'Luci', one of its constructor methods must be called -ResolutionErrors.dfy(3947,8): Error: when allocating an object of type 'Luci', one of its constructor methods must be called -ResolutionErrors.dfy(3971,44): Error: type of right argument to * (bool) must agree with the result type (int) -ResolutionErrors.dfy(4052,16): Error: array selection requires integer- or bitvector-based numeric indices (got bool for index 0) -ResolutionErrors.dfy(4050,13): Error: incorrect type for selection into array (got real) -ResolutionErrors.dfy(4051,11): Error: incorrect type for selection into seq (got real) -ResolutionErrors.dfy(4052,23): Error: array selection requires integer- or bitvector-based numeric indices (got real for index 1) -ResolutionErrors.dfy(4054,17): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4054,17): Error: incorrect type for selection into array (got real) -ResolutionErrors.dfy(4054,22): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4054,22): Error: incorrect type for selection into array (got real) -ResolutionErrors.dfy(4055,17): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4055,17): Error: incorrect type for selection into array (got real) -ResolutionErrors.dfy(4056,19): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4056,19): Error: incorrect type for selection into array (got real) -ResolutionErrors.dfy(4057,15): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4057,15): Error: incorrect type for selection into seq (got real) -ResolutionErrors.dfy(4057,20): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4057,20): Error: incorrect type for selection into seq (got real) -ResolutionErrors.dfy(4058,15): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4058,15): Error: incorrect type for selection into seq (got real) -ResolutionErrors.dfy(4059,17): Error: wrong number of indices for multi-selection -ResolutionErrors.dfy(4059,17): Error: incorrect type for selection into seq (got real) -650 resolution/type errors detected in ResolutionErrors.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy new file mode 100644 index 00000000000..8e1312b5fa5 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy @@ -0,0 +1,516 @@ +// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +module Misc { + //Should not verify, as ghost loops should not be allowed to diverge. + method GhostDivergentLoop() + { + var a := new int [2]; + a[0] := 1; + a[1] := -1; + ghost var i := 0; + while (i < 2) + decreases * // error: not allowed on a ghost loop + invariant i <= 2 + invariant (forall j :: 0 <= j && j < i ==> a[j] > 0) + { + i := 0; + } + assert a[1] != a[1]; // ...for then this would incorrectly verify + } + + method ManyIndices(a: array3, b: array, m: int, n: int) + { + // the following invalid expressions were once incorrectly resolved: + var x := a[m, n]; // error + var y := a[m]; // error + var z := b[m, n, m, n]; // error + } + + method SB(b: array2, s: int) returns (x: int, y: int) + requires b != null + { + while + { + case b[x,y] == s => + } + } + + // -------- name resolution + + class Global { + var X: int + function F(x: int): int { x } + static ghost function G(x: int): int { x } + method M(x: int) returns (r: int) + { + r := x + X; + } + static method N(x: int) returns (r: int) + { + r := x + X; // error: cannot access instance field X from static method + } + } + + method TestNameResolution0() { + var z: int; + z := Global.X; // error: X is an instance field + z := F(2); // error: cannot resolve F + z := Global.F(2); // error: invocation of instance function requires an instance + z := G(2); // error: cannot resolve G + z := Global.G(2); + z := M(2); // error: cannot resolve M + z := Global.M(2); // error: call to instance method requires an instance + z := N(1); // error: cannot resolve N + z := Global.N(1); + + z := z(5); // error: using local as if it were a function + z := Global.z; // error: class Global does not have a member z + + var Global: Global; // a local variable with the name 'Global' + z := Global.X; // this means the instance field X of the object stored in the local variable 'Global' + var gg: Global := null; + var y := gg.G(5); + y := gg.N(5); + } + + datatype Abc = Abel | Benny | Cecilia(y: int) | David(x: int) | Eleanor + datatype Xyz = Alberta | Benny | Constantine(y: int) | David(x: int) + datatype Rst = David(x: int, y: int) + + ghost function Tuv(arg0: Abc, arg1: bool): int { 10 } + + class EE { + var Eleanor: bool + method TestNameResolution1() { + var a0 := Abel; + var a1 := Alberta; + var b0 := Benny; // error: there's more than one constructor with the name Benny; needs qualification + var b1 := Abc.Benny; + var b2 := Xyz.Benny; + var Benny := 15; // introduce a local variable with the name 'Benny' + var b3 := Benny; + var d0 := David(20); // error: constructor name David is ambiguous + var d1 := David; // error: constructor name David is ambiguous (never mind that the signature does + // not match either of them) + var d2 := David(20, 40); // error: constructor name Davis is ambiguous (never mind that the given + // parameters match the signature of only one of those constructors) + var d3 := Abc.David(20, 40); // error: wrong number of parameters + var d4 := Rst.David(20, 40); + var e := Eleanor; // this resolves to the field, not the Abc datatype constructor + assert Tuv(Abc.Eleanor, e) == 10; + } + } +} + +// --------------- ghost tests ------------------------------------- + +module HereAreMoreGhostTests { + datatype GhostDt = + Nil(ghost extraInfo: int) | + Cons(data: int, tail: GhostDt, ghost moreInfo: int) + + class GhostTests { + method M(dt: GhostDt) returns (r: int) { + ghost var g := 5; + r := g; // error: RHS is ghost, LHS is not + r := F(18, g); // error: RHS is a ghost and will not be available at run time + r := G(20, g); // it's fine to pass a ghost as a parameter to a non-ghost, because + // only the ghost goes away during compilation + r := N(22, g); // ditto + r := N(g, 22); // error: passing in 'g' as non-ghost parameter + r := P(24, 22); // error: 'P' is ghost, but its result is assigned to a non-ghost + + match (dt) { + case Nil(gg) => + case Cons(dd, tt, gg) => + r := G(dd, dd); // fine + r := G(dd, gg); // fine + r := G(gg, gg); // error: cannot pass ghost 'gg' as non-ghost parameter to 'G' + } + var dd; + dd := GhostDt.Nil(g); // fine + dd := GhostDt.Cons(g, dt, 2); // error: cannot pass 'g' as non-ghost parameter + ghost var dtg := GhostDt.Cons(g, dt, 2); // fine, since result is ghost + } + ghost function F(x: int, y: int): int { + y + } + function G(x: int, ghost y: int): int { + y // error: cannot return a ghost from a non-ghost function + } + function H(dt: GhostDt): int { + match dt + case Nil(gg) => gg // error: cannot return a ghost from a non-ghost function + case Cons(dd, tt, gg) => dd + gg // error: ditto + } + method N(x: int, ghost y: int) returns (r: int) { + r := x; + } + ghost method P(x: int, y: int) returns (r: int) { + ghost var g := 5; + r := y; // allowed, since the entire method is ghost + r := r + g; // fine, for the same reason + r := N(20, 20); // error: call to non-ghost method from ghost method is not okay + } + ghost method BreaksAreFineHere(t: int) + { + var n := 0; + ghost var k := 0; + while (true) + invariant n <= 112 + decreases 112 - n + { + label MyStructure: { + if (k % 17 == 0) { break MyStructure; } // this is fine, because it's a ghost method + k := k + 1; + } + label MyOtherStructure: + if (k % 17 == 0) { + break MyOtherStructure; + } else { + k := k + 1; + } + + if (n == 112) { + break; + } else if (n == t) { + return; + } + n := n + 1; + } + } + method BreakMayNotBeFineHere(ghost t: int) + { + var n := 0; + ghost var k := 0; + var p := 0; + while (true) + invariant n <= 112 + decreases 112 - n + { + label MyStructure: { + k := k + 1; + } + label MyOtherStructure: + if (k % 17 == 0) { + break MyOtherStructure; // this break is fine + } else { + k := k + 1; + } + + var dontKnow; + if (n == 112) { + ghost var m := 0; + label LoopLabel0: + label LoopLabel1: + while (m < 200) { + if (m % 103 == 0) { + if { + case true => break; // fine, since this breaks out of the enclosing ghost loop + case true => break LoopLabel0; // fine + case true => break LoopLabel1; // fine + } + } else if (m % 101 == 0) { + } + m := m + 3; + } + break; + } else if (dontKnow == 708) { + var q := 0; + while (q < 1) { + label IfNest: + if (p == 67) { + break break; // fine, since this is not a ghost context + } + q := q + 1; + } + } else if (n == t) { + } + n := n + 1; + p := p + 1; + } + } + method BreakMayNotBeFineHere_Ghost(ghost t: int) + { + var n := 0; + ghost var k := 0; + var p := 0; + while (true) + invariant n <= 112 + decreases 112 - n + { + label MyStructure: { + if (k % 17 == 0) { break MyStructure; } // error: break from ghost to non-ghost point + k := k + 1; + } + label MyOtherStructure: + if (k % 17 == 0) { + break MyOtherStructure; // this break is fine + } else { + k := k + 1; + } + + var dontKnow; + if (n == 112) { + ghost var m := 0; + label LoopLabel0: + label LoopLabel1: + while (m < 200) { + if (m % 103 == 0) { + if { + case true => break; // fine, since this breaks out of the enclosing ghost loop + case true => break LoopLabel0; // fine + case true => break LoopLabel1; // fine + } + } else if (m % 101 == 0) { + break break; // error: break out of non-ghost loop from ghost context + } + m := m + 3; + } + break; + } else if (dontKnow == 708) { + var q := 0; + while (q < 1) { + label IfNest: + if (p == 67) { + break break; // fine, since this is not a ghost context + } else if (*) { + break break; // fine, since this is not a ghost context + } else if (k == 67) { + break break; // error, because this is a ghost context + } + q := q + 1; + } + } else if (n == t) { + return; // error: this is a ghost context trying to return from a non-ghost method + } + n := n + 1; + p := p + 1; + } + } + } +} // HereAreMoreGhostTests + +module MiscMore { + method DuplicateLabels(n: int) { + var x; + if (n < 7) { + label DuplicateLabel: x := x + 1; + } else { + label DuplicateLabel: x := x + 1; + } + label DuplicateLabel: x := x + 1; + label DuplicateLabel': { + label AnotherLabel: + label DuplicateLabel': // error: duplicate label (shadowed by enclosing label) + label OneMoreTime: + x := x + 1; + } + label DuplicateLabel'': + label DuplicateLabel'': // error: duplicate label (shadowed by enclosing label) + x := x + 1; + label DuplicateLabel'': x := x + 1; // error: duplicate label (shadowed by dominating label) + } + + // --------------- constructors ------------------------------------- + // Note, more tests in module ClassConstructorTests below + + class ClassWithConstructor { + var y: int + method NotTheOne() { } + constructor InitA() { } + constructor InitB() modifies this { y := 20; } // error: don't use "this" in modifies of constructor + } + + method ConstructorTests() + { + var c := new ClassWithConstructor.InitB(); + c.InitB(); // error: not allowed to call constructors except during allocation + } + + // ------------------- datatype destructors --------------------------------------- + + datatype DTD_List = DTD_Nil | DTD_Cons(Car: int, Cdr: DTD_List, ghost g: int) + + method DatatypeDestructors(d: DTD_List) { + if { + case d.DTD_Nil? => + assert d == DTD_Nil; + case d.DTD_Cons? => + var hd := d.Car; + var tl := d.Cdr; + assert hd == d.Cdr; // type error + assert tl == d.Car; // type error + assert d.DTD_Cons? == d.Car; // type error + assert d == DTD_Cons(hd, tl, 5); + ghost var g0 := d.g; // fine + } + } +} // MiscMore + +// ------------------- print statements --------------------------------------- + +module GhostPrintAttempts { + method PrintOnlyNonGhosts(a: int, ghost b: int) + { + print "a: ", a, "\n"; + print "b: ", b, "\n"; // error: print statement cannot take ghosts + } +} + +// ------------------- auto-added type arguments ------------------------------ + +module MiscEvenMore { + class GenericClass { var data: T } + + method MG0(a: GenericClass, b: GenericClass) + requires a != null && b != null + modifies a + { + a.data := b.data; // allowed, since both a and b get the same auto type argument + } + + method G_Caller() + { + var x := new GenericClass; + MG0(x, x); // fine + var y := new GenericClass; + MG0(x, y); // also fine (and now y's type argument is constrained to be that of x's) + var z := new GenericClass; + var w := new GenericClass; + MG0(x, w); // this has the effect of making x's and y's type GenericClass + + y.data := z.data; // error: bool vs int + assert x.data == 5; // error: bool vs int + } + + datatype GList<+T> = GNil | GCons(hd: T, tl: GList) + + method MG1(l: GList, n: nat) + { + if (n != 0) { + MG1(l, n-1); + MG1(GCons(12, GCons(20, GNil)), n-1); + } + var t := GCons(100, GNil); // error: types don't match up (List<_T0> versus List) + t := GCons(120, l); // error: types don't match up (List<_T0> versus List) + } + + // ------------------- calc statements ------------------------------ + + method TestCalc(m: int, n: int, a: bool, b: bool) + { + calc { + a + b; // error: invalid line + n + m; + } + calc { + a && b; + n + m; // error: all lines must have the same type + } + calc ==> { + n + m; // error: ==> operator requires boolean lines + n + m + 1; + n + m + 2; + } + calc { + n + m; + n + m + 1; + ==> n + m + 2; // error: ==> operator requires boolean lines + } + } +} // MiscEvenMore + +module MyOwnModule { + class SideEffectChecks { + ghost var ycalc: int + + ghost method Mod(a: int) + modifies this + ensures ycalc == a + { + ycalc := a; + } + + ghost method Bad() + modifies this + ensures 0 == 1 + { + var x: int; + calc { + 0; + { Mod(0); } // error: methods with side-effects are not allowed + ycalc; + { ycalc := 1; } // error: heap updates are not allowed + 1; + { x := 1; } // error: updates to locals defined outside of the hint are not allowed + x; + { + var x: int; + x := 1; // this is OK + } + 1; + } + } + } +} + +// ------------------- nameless constructors ------------------------------ + +module MiscAgain { + class Y { + var data: int + constructor (x: int) + { + data := x; + } + constructor (y: bool) // error: duplicate constructor name + { + } + method Test() { + var i := new Y(5); + i := new Y(7); + + var s := new Luci.Init(5); + s := new Luci.FromArray(null); + s := new Luci(false); + s := new Luci(true); + + var l := new Lamb; + l := new Lamb(); // error: there is no default constructor + } + } + + class Luci { + constructor Init(y: int) { } + constructor (nameless: bool) { } + constructor FromArray(a: array) { } + method M() { } + } + + class Lamb { + method Jess() { } + method Gwen() { } + } + + // ------------------- assign-such-that and ghosts ------------------------------ + + method AssignSuchThatFromGhost() + { + var x: int; + ghost var g: int; + + x := *; + assume x == g; // this mix of ghosts and non-ghosts is cool (but, of course, + // the compiler will complain) + + x :| assume x == g; // this is cool, since it's an assume (but, of course, the + // compiler will complain) + + x :| x == 5; + g :| g <= g; + g :| assume g < g; // the compiler will complain here, despite the LHS being + // ghost -- and rightly so, since an assume is used + } +} // MiscAgain diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy.expect new file mode 100644 index 00000000000..28211560fe1 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors0.dfy.expect @@ -0,0 +1,70 @@ +ResolutionErrors0.dfy(323,33): Warning: constructors no longer need 'this' to be listed in modifies clauses +ResolutionErrors0.dfy(13,18): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating +ResolutionErrors0.dfy(25,13): Error: array selection requires an array2 (got array3) +ResolutionErrors0.dfy(27,13): Error: array selection requires an array4 (got array) +ResolutionErrors0.dfy(26,14): Error: element selection requires a sequence, array, multiset, or map (got array3) +ResolutionErrors0.dfy(57,16): Error: accessing member 'X' requires an instance expression +ResolutionErrors0.dfy(58,9): Error: unresolved identifier: F +ResolutionErrors0.dfy(59,16): Error: accessing member 'F' requires an instance expression +ResolutionErrors0.dfy(60,9): Error: unresolved identifier: G +ResolutionErrors0.dfy(62,9): Error: unresolved identifier: M +ResolutionErrors0.dfy(63,16): Error: accessing member 'M' requires an instance expression +ResolutionErrors0.dfy(64,9): Error: unresolved identifier: N +ResolutionErrors0.dfy(67,10): Error: non-function expression (of type int) is called with parameters +ResolutionErrors0.dfy(68,16): Error: member 'z' does not exist in non-null type 'Global' +ResolutionErrors0.dfy(51,15): Error: 'this' is not allowed in a 'static' context +ResolutionErrors0.dfy(88,16): Error: the name 'Benny' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.Benny') +ResolutionErrors0.dfy(93,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') +ResolutionErrors0.dfy(94,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') +ResolutionErrors0.dfy(96,16): Error: the name 'David' denotes a datatype constructor, but does not do so uniquely; add an explicit qualification (for example, 'Abc.David') +ResolutionErrors0.dfy(98,20): Error: wrong number of arguments (got 2, but datatype constructor 'David' expects 1: (x: int)) +ResolutionErrors0.dfy(116,11): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(117,11): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors0.dfy(121,13): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(122,12): Error: actual out-parameter is required to be a ghost variable +ResolutionErrors0.dfy(129,17): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(133,25): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(140,6): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(144,23): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(145,37): Error: ghost variables such as gg are allowed only in specification contexts. gg was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(154,12): Error: only ghost methods can be called from this context +ResolutionErrors0.dfy(244,29): Error: ghost-context break statement is not allowed to break out of non-ghost structure +ResolutionErrors0.dfy(267,14): Error: ghost-context break statement is not allowed to break out of non-ghost loop +ResolutionErrors0.dfy(281,14): Error: ghost-context break statement is not allowed to break out of non-ghost loop +ResolutionErrors0.dfy(286,10): Error: return statement is not allowed in this context (because it is guarded by a specification-only expression) +ResolutionErrors0.dfy(306,12): Error: label shadows an enclosing label +ResolutionErrors0.dfy(311,10): Error: duplicate label +ResolutionErrors0.dfy(313,10): Error: label shadows a dominating label +ResolutionErrors0.dfy(329,11): Error: a constructor is allowed to be called only when an object is being allocated +ResolutionErrors0.dfy(345,27): Error: arguments must have comparable types (got bool and int) +ResolutionErrors0.dfy(343,18): Error: arguments must have comparable types (got int and DTD_List) +ResolutionErrors0.dfy(344,18): Error: arguments must have comparable types (got DTD_List and int) +ResolutionErrors0.dfy(358,17): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. +ResolutionErrors0.dfy(382,7): Error: incorrect argument type at index 1 for method in-parameter 'b' (expected GenericClass, found GenericClass) (non-variant type parameter would require int = bool) +ResolutionErrors0.dfy(396,13): Error: incorrect argument type at index 0 for datatype constructor parameter 'hd' (expected _T0, found int) +ResolutionErrors0.dfy(397,9): Error: incorrect argument type at index 0 for datatype constructor parameter 'hd' (expected _T0, found int) +ResolutionErrors0.dfy(406,8): Error: all lines in a calculation must have the same type (got int after bool) +ResolutionErrors0.dfy(410,8): Error: all lines in a calculation must have the same type (got int after bool) +ResolutionErrors0.dfy(413,8): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(413,8): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(414,12): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(414,12): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(415,12): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(415,12): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(413,8): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(413,8): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(419,12): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(419,12): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(418,8): Error: first argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(418,8): Error: second argument to ==> must be of type bool (instead got int) +ResolutionErrors0.dfy(405,8): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) +ResolutionErrors0.dfy(405,8): Error: arguments must have comparable types (got bool and int) +ResolutionErrors0.dfy(405,8): Error: arguments must have comparable types (got bool and int) +ResolutionErrors0.dfy(409,8): Error: arguments must have comparable types (got bool and int) +ResolutionErrors0.dfy(409,8): Error: arguments must have comparable types (got bool and int) +ResolutionErrors0.dfy(443,13): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors0.dfy(445,10): Error: a hint is not allowed to make heap updates +ResolutionErrors0.dfy(447,10): Error: a hint is not allowed to update a variable it doesn't declare +ResolutionErrors0.dfy(468,4): Error: More than one anonymous constructor +ResolutionErrors0.dfy(481,15): Error: class Lamb does not have an anonymous constructor +68 resolution/type errors detected in ResolutionErrors0.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy new file mode 100644 index 00000000000..38f344cf633 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy @@ -0,0 +1,182 @@ +// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ------------------------ inferred type arguments ---------------------------- + +// Put the following tests in a separate module, so that the method bodies will +// be type checked even if there are resolution errors in other modules. + +module NoTypeArgs0 { + datatype List<+T> = Nil | Cons(T, List) + datatype Tree<+A,+B> = Leaf(A, B) | Node(Tree, Tree) + + method DoAPrefix0(xs: List) returns (ys: List) + { + ys := xs; + } + + method DoAPrefix1(xs: List) returns (ys: List) + { + ys := xs; // error: List cannot be assign to a List + } + + method DoAPrefix2(xs: List) returns (ys: List) + { + ys := xs; // error: List cannot be assign to a List + } + + ghost function FTree0(t: Tree): Tree + { + match t + case Leaf(_,_) => t + case Node(x, y) => x + } + + ghost function FTree1(t: Tree): Tree + { + match t + case Leaf(_,_) => t + case Node(x, y) => y // error: y does not have the right type + } + + ghost function FTree2(t: Tree): Tree + { + t + } +} + +module NoTypeArgs1 { + datatype Tree = Leaf(A, B) | Node(Tree, Tree) + + ghost function FTree3(t: Tree): Tree // error: type of 't' does not have enough type parameters + { + t + } +} + +// ----------- let-such-that expressions ------------------------ + +module MiscMisc { + method LetSuchThat(ghost z: int, n: nat) + { + var x: int; + x := var y :| y < 0; y; // fine for the resolver (but would give a verification error for not being deterministic) + + x := var w :| w == 2*w; w; // fine (even for the verifier, this one) + x := var w := 2*w; w; // error: the 'w' in the RHS of the assignment is not in scope + ghost var xg := var w :| w == 2*w; w; + } +} + +// ------------ quantified variables whose types are not inferred ---------- + +module NonInferredType { + ghost predicate P(x: T) + + method InferredType(x: int) + { + var t; + assume forall z :: P(z) && z == t; + assume t == x; // this statement determines the type of t and z + } + + method NonInferredType(x: int) + { + var t; // error: the type of t is not determined + assume forall z :: P(z) && z == t; // error: the type of z is not determined + } +} + +// ------------ Here are some tests that lemma contexts don't allocate objects ------------- + +module GhostAllocationTests { + class G { } + iterator GIter() { } + class H { constructor () } + lemma GhostNew0() + ensures exists o: G :: fresh(o) + { + var p := new G; // error: lemma context is not allowed to allocate state + p := new G; // error: ditto + } + + method GhostNew1(n: nat, ghost g: int) returns (t: G, z: int) + { + if n < 0 { + z, t := 5, new G; // fine + } + if n < g { + var tt := new H(); // error: 'new' not allowed in ghost contexts + } + } + + method GhostNew2(ghost b: bool) + { + if (b) { + var y := new GIter(); // error: 'new' not allowed in ghost contexts (and a non-ghost method is not allowed to be called here either) + } + } +} +module MoreGhostAllocationTests { + class G { } + method GhostNew3(n: nat) { + var g := new G; + calc { + 5; + { var y := new G; } // error: 'new' not allowed in lemma contexts + 2 + 3; + } + } + ghost method GhostNew4(g: G) + modifies g + { + } +} + +module NewForallAssign { + class G { } + method NewForallTest(n: nat) { + var a := new G[n]; + forall i | 0 <= i < n { + a[i] := new G; // error: 'new' is currently not supported in forall statements + } } +} +module NewForallProof { + class G { } + method NewForallTest(n: nat) { var a := new G[n]; + forall i | 0 <= i < n ensures true { // this makes the whole 'forall' statement into a ghost statement + a[i] := new G; // error: proof-forall cannot update state (and 'new' not allowed in ghost contexts, but that's checked at a later stage) + } } +} + +// ------------------------- underspecified types ------------------------------ + +module UnderspecifiedTypes { + method M(S: set) { + var n, p, T0 :| 12 <= n && n in T0 && 10 <= p && p in T0 && T0 <= S && p % 2 != n % 2; + var T1 :| 12 in T1 && T1 <= S; + var T2 :| T2 <= S && 12 in T2; + var T3 :| 120 in T3; // error: underspecified type + var T3'0: set :| 120 in T3'0; + var T3'1: multiset :| 120 in T3'1; + var T3'2: map :| 120 in T3'2; + var T3'3: seq :| 120 in T3'3; + var T3'4: bool :| 120 in T3'4; // error: second argument to 'in' cannot be bool + var T4 :| T4 <= S; + } +} + +// ------------------------- lemmas ------------------------------ + +module MiscLemma { + class L { } + + // a lemma is allowed to have out-parameters, but not a modifies clause + lemma MyLemma(x: int, l: L) returns (y: int) + requires 0 <= x + modifies l + ensures 0 <= y + { + y := x; + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy.expect new file mode 100644 index 00000000000..9e74df5f970 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors1.dfy.expect @@ -0,0 +1,21 @@ +ResolutionErrors1.dfy(20,7): Error: RHS (of type List) not assignable to LHS (of type List) (covariant type parameter would require A <: B) +ResolutionErrors1.dfy(25,7): Error: RHS (of type List) not assignable to LHS (of type List) (covariant type parameter would require A <: B) +ResolutionErrors1.dfy(39,23): Error: type of case bodies do not agree (found Tree<_T1, _T0>, previous types Tree<_T0, _T1>) (covariant type parameter 0 would require _T1 <: _T0) +ResolutionErrors1.dfy(51,30): Error: Wrong number of type arguments (0 instead of 2) passed to datatype: Tree +ResolutionErrors1.dfy(66,20): Error: unresolved identifier: w +ResolutionErrors1.dfy(85,8): Error: the type of this local variable is underspecified +ResolutionErrors1.dfy(86,25): Error: the type of this variable is underspecified +ResolutionErrors1.dfy(86,23): Error: type parameter 'T' (inferred to be '?') in the function call to 'P' could not be determined +ResolutionErrors1.dfy(86,18): Error: type of bound variable 'z' could not be determined; please specify the type explicitly +ResolutionErrors1.dfy(99,13): Error: a lemma is not allowed to use 'new' +ResolutionErrors1.dfy(100,9): Error: a lemma is not allowed to use 'new' +ResolutionErrors1.dfy(109,16): Error: only ghost methods can be called from this context +ResolutionErrors1.dfy(116,15): Error: only ghost methods can be called from this context +ResolutionErrors1.dfy(126,17): Error: a hint is not allowed to use 'new' +ResolutionErrors1.dfy(141,14): Error: new allocation not supported in aggregate assignments +ResolutionErrors1.dfy(148,14): Error: a forall statement is not allowed to use 'new' +ResolutionErrors1.dfy(148,11): Error: assignment to array element is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors1.dfy(164,26): Error: second argument to "in" must be a set, multiset, or sequence with elements of type int, or a map with domain int (instead got bool) +ResolutionErrors1.dfy(159,18): Error: second argument to "in" must be a set, multiset, or sequence with elements of type int, or a map with domain int (instead got ?) +ResolutionErrors1.dfy(177,13): Error: lemmas are not allowed to have modifies clauses +20 resolution/type errors detected in ResolutionErrors1.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy new file mode 100644 index 00000000000..7ebfece166b --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy @@ -0,0 +1,525 @@ +// RUN: %exits-with 2 %verify "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ------------------------- statements in expressions ------------------------------ + +module StatementsInExpressions { + class MyClass { + ghost method SideEffect() + modifies this + { + } + + method NonGhostMethod() + { + } + + ghost function F(): int + { + calc { + 6; + { assert 6 < 8; } + { var x := 8; + while x != 0 + decreases * // error: cannot use 'decreases *' here + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + 6; + } + 5 + } + + var MyField: int + ghost var MyGhostField: int + + method N() + { + var y := + calc { + 6; + { assert 6 < 8; } + { var x := 8; + while x != 0 + decreases * // error: cannot use 'decreases *' here + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + 6; + } + 5; + } + } +} + +module StmtExprOutParams { + + lemma MyLemma() + + lemma OutParamLemma() returns (y: int) + + ghost function UseLemma(): int + { + MyLemma(); + OutParamLemma(); // error: has out-parameters + 10 + } +} + +module GhostLetExpr { + method M() { + ghost var y; + var x; + var g := G(x, y); + ghost var h := ghost var ta := F(); 5; + var j; j := var tb := F(); 5; // fine (tb is ghost, j is not, but RHS body doesn't depend on tb) + assert h == j; + } + + ghost function F(): int { + 5 + } + + function G(x: int, ghost y: int): int { + assert y == x; + y // error: not allowed in non-ghost context + } + + datatype Dt = MyRecord(a: int, ghost b: int) + + method P(dt: Dt) { + match dt { + case MyRecord(aa, bb) => + ghost var z := aa + F(); + ghost var t0 := var y := z; z + 3; + ghost var t1 := ghost var y := z + bb; y + z + 3; + var t2; t2 := ghost var y := z; y + 3; // t2 is not ghost - error + } + } + + function FM(e: bool): int { + if e then + G(5, F()) + else + var xyz := F(); // fine, because 'xyz' becomes ghost automatically + G(5, xyz) + } +} + +module ObjectType { + type B + datatype Dt = Blue | Green + codatatype CoDt = Cons(int, CoDt) + class MyClass { } + + method M(zz: array, j: int, b: B, co: CoDt, g: G) returns (o: object) + requires zz != null && 0 <= j < zz.Length + { + o := b; // error + o := 17; // error + o := zz[j]; // error + o := null; + o := zz; + o := new MyClass; + o := o; + o := g; // error + o := Blue; // error + o := co; // error + } +} + +// ------------------ modify statment --------------------------- + +module MiscModify { + class ModifyStatementClass { + var x: int + ghost var g: int + method M() { + modify x; // error: type error + } + } +} + +module MiscModifiesGhost { + class ModifyStatementClass { + var x: int + ghost var g: int + ghost method G0() + modifies `g + modifies `x // error: non-ghost field mentioned in ghost context + } +} + +module ModifyStatementClass_More { + class C { + var x: int + ghost var g: int + + ghost method G0() + modifies `g + { + modify `g; + modify `x; // error: non-ghost field mentioned in ghost context + } + + method G1() + modifies this + { + modify `x; + if g < 100 { + // we are now in a ghost context + modify `x; // error: non-ghost field mentioned in ghost context + } + } + + method G2(y: nat) + modifies this + { + if g < 100 { + // we're now in a ghost context + var n := 0; + while n < y + modifies `x // error: non-ghost field mentioned in ghost context + { + if * { + g := g + 1; // if we got as far as verification, this would be flagged as an error too + } + n := n + 1; + } + } + modify `x; // fine + + ghost var i := 0; + while i < y + modifies `x // error: non-ghost field mentioned in ghost context + { + i := i + 1; + } + } + } +} + +module LhsLvalue { + method M() + { + var mySeq: seq; + var a := new int[78]; + var b := new int[100, 200]; + var c := new MyRecord[29]; + + mySeq[0] := 5; // error: cannot assign to a sequence element + mySeq[0] := MyMethod(); // error: ditto + a[0] := 5; + a[0] := MyMethod(); + b[20, 18] := 5; + b[20, 18] := MyMethod(); + c[25].x := 5; // error: cannot assign to a destructor + c[25].x := MyMethod(); // error: ditto + mySeq[0..4] := 5; // error: cannot assign to a range + mySeq[0..4] := MyMethod(); // error: ditto + a[0..4] := 5; // error: cannot assign to a range + a[0..4] := MyMethod(); // error: ditto + } + + datatype MyRecord = Make(x: int, y: int) + + method MyMethod() returns (w: int) +} + +// ------------------- dirty loops ------------------- + +module MiscEtc { + method DirtyM(S: set) { + forall s | s in S ensures s < 0 + assert s < 0; // error: s is unresolved + } + + // ------------------- tuples ------------------- + + method TupleResolution(x: int, y: int, r: real) + { + var unit: () := (); + var expr: int := (x); + var pair: (int,int) := (x, x); + var triple: (int,int,int) := (y, x, x); + var badTriple: (int,real,int) := (y, x, r); // error: parameters 1 and 2 have the wrong types + var quadruple: (int,real,int,real) := (y, r, x); // error: trying to use a triple as a quadruple + + assert unit == (); + assert pair.0 == pair.1; + assert triple.2 == x; + + assert triple.2; // error: 2 has type int, not the expected bool + assert triple.3 == pair.x; // error(s): 3 and x are not destructors + + var k0 := (5, (true, 2, 3.14)); + var k1 := (((false, 10, 2.7)), 100, 120); + if k0.1 == k1.0 { + assert false; + } else if k0.1.1 < k1.0.1 { + assert k1.2 == 120; + } + + // int and (int) are the same type (i.e., there are no 1-tuples) + var pp: (int) := x; + var qq: int := pp; + } + + // ------------------- conversions ------------------- + + method TypeConversions(m: nat, i: int, r: real) returns (n: nat, j: int, s: real) + { + n := r as int; + j := r as int; + s := m as real; // nat->real is allowed, just like int->real is + s := i as real; + s := i as real / 2; // error: division expects two reals + s := 15 % s; // error: modulus is not defined for reals + + s := (2.0 / 1.7) + (r / s) - (--r) * -12.3; + + s := s as real; // fine (identity transform) + j := j as int; // fine (identity transform) + j := n as int; // fine (identity transform) + } +} + +// --- filling in type arguments and checking that there aren't too many --- + +module TypeArgumentCount { + class C { + var f: T + } + + method R0(a: array3, c: C) + + method R1() + { + var a: array3; + var c: C; + } + + method R2() + { + var a: array3; // error: too many type arguments + var c: C; // error: too many type arguments + } +} + +// --- Type synonyms --- + +module BadTypeSynonyms { + datatype List = Nil | Cons(T, List) + type BadSyn0 = List // error: must have at least one type parameter + type BadSyn1 = badName // error: badName does not denote a type + type BadSyn2 = List // error: X does not denote a type + type BadSyn2 = int // error: repeated name +} + +// --- cycles --- + +module CycleError0 { + type A = A // error: cycle: A -> A +} + +module CycleError1 { + type A = B // error: cycle: A -> B -> A + type B = A +} + +module CycleError2 { + type A = B // error: cycle: A -> B -> A + type B = set +} + +module CycleErrors3 { + type A = (B, D) + type B = C + class C { // error: since A is not auto-init, class C needs a constructor + var a: A // this is fine + } + datatype D = Make(A, B, C) // warning: D is empty +} + +module CycleError4 { + type A = B // error: cycle: A -> B -> A + type B = C + class C { } +} + +module CycleError5 { + type A = B // error: cycle: A -> B -> A + type B = Dt + datatype Dt = Make(T) +} + +module CycleError6 { + type A = Dt> // error: cycle A -> A + datatype Dt = Make(T) +} + +// --- attributes in top-level declarations --- + +module MiscIterator { + iterator {:myAttribute x} Iter() { // error: x does not refer to anything + } + + class {:myAttribute x} C { // error: x does not refer to anything + } + + datatype {:myAttribute x} Dt = Blue // error: x does not refer to anything + + type {:myAttribute x} Something // error: x does not refer to anything + + type {:myAttribute x} Synonym = int // error: x does not refer to anything +} + +module {:myAttribute x} Modulette { // error: x does not refer to anything +} + +// --- abstract types with type parameters --- + +module OpaqueTypes0 { + type P + method M(p: P) returns (q: P) // error: wrong param count + { + q := p; + } +} + +module OpaqueTypes1 { + type P + + method M0(p: P) returns (q: P) + { + q := p; + var m: P; // error: BX undefined + } + + method M1(p: P) returns (q: P) // type parameter of q's type inferred + { + q := p; + } + + method M2(p: P) returns (q: P) + { + q := p; // error: cannot assign P to P + } + + method M3(p: P) returns (q: P) + { + q := p; // error: cannot assign P to P + } + + method M4() returns (p: P, q: P) + { + q := p; // error: cannot assign P to P + p := q; // error: cannot assign P to P + } + + method EqualityTests(p: P, q: P, r: P) + { + assert p != r; // error: types must be the same in order to do compare + assert q != r; // error: types must be the same in order to do compare + assert p != q; // error: types must be the same in order to do compare + } +} + +// ----- new trait ------------------------------------------- + +module MiscTrait { + trait J { } + type JJ = J + method TraitSynonym() + { + var x := new JJ; // error: new cannot be applied to a trait + } +} + +// ----- set comprehensions where the term type is finite ----- + +module ObjectSetComprehensionsNever { + // the following set comprehensions are known to be finite + ghost function A() : set { set o : object | true :: o } // error: a function is not allowed to depend on the allocated state + function B() : set { set o : object | true :: o } // error: a function is not allowed to depend on the allocated state +} + +module ObjectSetComprehensionsSometimes { + // outside functions, the comprehension is permitted, but it cannot be compiled + lemma C() { var x; x := set o : object | true :: o; } + + method D() { var x; x := set o : object | true :: o; } // error: not (easily) compilable, so this is allowed only in ghost contexts +} + +// ------ regression test for type checking of integer division ----- + +module MiscTests { + method IntegerDivision(s: set) + { + var t := s / s; // error: / cannot be used with sets + } + + // ----- decreases * tests ---- + + method NonTermination_A() + { + NonTermination_B(); // error: to call a non-terminating method, the caller must be marked 'decreases *' + } + + method NonTermination_B() + decreases * + { + while true + decreases * + { + } + } + + method NonTermination_C() + { + while true + decreases * // error: to use an infinite loop, the enclosing method must be marked 'decreases *' + { + } + } + + method NonTermination_D() + decreases * + { + var n := 0; + while n < 100 // note, no 'decreases *' here, even if the nested loop may fail to terminate + { + while * + decreases * + { + } + n := n + 1; + } + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy.expect new file mode 100644 index 00000000000..00c24866355 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors2.dfy.expect @@ -0,0 +1,73 @@ +ResolutionErrors2.dfy(24,22): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating +ResolutionErrors2.dfy(57,22): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating +ResolutionErrors2.dfy(90,17): Error: the method returns 1 value but is assigned to 0 variable (all return values must be assigned) +ResolutionErrors2.dfy(111,4): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. +ResolutionErrors2.dfy(122,40): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. +ResolutionErrors2.dfy(144,6): Error: RHS (of type B) not assignable to LHS (of type object) +ResolutionErrors2.dfy(146,6): Error: RHS (of type B) not assignable to LHS (of type object) +ResolutionErrors2.dfy(151,6): Error: RHS (of type G) not assignable to LHS (of type object) +ResolutionErrors2.dfy(152,6): Error: RHS (of type Dt) not assignable to LHS (of type object) +ResolutionErrors2.dfy(153,6): Error: RHS (of type CoDt) not assignable to LHS (of type object) +ResolutionErrors2.dfy(145,6): Error: RHS (of type int) not assignable to LHS (of type object) +ResolutionErrors2.dfy(164,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int) +ResolutionErrors2.dfy(175,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors2.dfy(188,14): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors2.dfy(197,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors2.dfy(208,20): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors2.dfy(220,18): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors2.dfy(242,10): Error: LHS of assignment must denote a mutable field +ResolutionErrors2.dfy(243,10): Error: LHS of assignment must denote a mutable field +ResolutionErrors2.dfy(244,9): Error: cannot assign to a range of array elements (try the 'forall' statement) +ResolutionErrors2.dfy(245,9): Error: cannot assign to a range of array elements (try the 'forall' statement) +ResolutionErrors2.dfy(246,5): Error: cannot assign to a range of array elements (try the 'forall' statement) +ResolutionErrors2.dfy(247,5): Error: cannot assign to a range of array elements (try the 'forall' statement) +ResolutionErrors2.dfy(236,4): Error: LHS of array assignment must denote an array element (found seq) +ResolutionErrors2.dfy(237,4): Error: LHS of array assignment must denote an array element (found seq) +ResolutionErrors2.dfy(260,11): Error: unresolved identifier: s +ResolutionErrors2.dfy(279,18): Error: member '3' does not exist in datatype '_tuple#3' +ResolutionErrors2.dfy(279,28): Error: member 'x' does not exist in datatype '_tuple#2' +ResolutionErrors2.dfy(278,18): Error: condition is expected to be of type bool, but is int +ResolutionErrors2.dfy(272,39): Error: RHS (of type (int, real, int)) not assignable to LHS (of type (int, real, int, real)) +ResolutionErrors2.dfy(271,34): Error: RHS (of type (int, int, real)) not assignable to LHS (of type (int, real, int)) (covariant type parameter 1 would require int <: real) +ResolutionErrors2.dfy(271,34): Error: RHS (of type (int, int, real)) not assignable to LHS (of type (int, real, int)) (covariant type parameter 2 would require real <: int) +ResolutionErrors2.dfy(303,12): Error: type of left argument to % (int) must agree with the result type (real) +ResolutionErrors2.dfy(303,12): Error: arguments to % must be integer-numeric or bitvector types (got real) +ResolutionErrors2.dfy(302,19): Error: type of right argument to / (int) must agree with the result type (real) +ResolutionErrors2.dfy(330,11): Error: Wrong number of type arguments (2 instead of 1) passed to non-null type: array3 +ResolutionErrors2.dfy(331,11): Error: Wrong number of type arguments (2 instead of 1) passed to non-null type: C +ResolutionErrors2.dfy(342,7): Error: duplicate name of top-level declaration: BadSyn2 +ResolutionErrors2.dfy(341,7): Related location +ResolutionErrors2.dfy(339,17): Error: Wrong number of type arguments (0 instead of 1) passed to datatype: List +ResolutionErrors2.dfy(340,17): Error: Type or type parameter is not declared in this scope: badName (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors2.dfy(341,22): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors2.dfy(348,7): Error: type-synonym cycle: A -> A +ResolutionErrors2.dfy(352,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A +ResolutionErrors2.dfy(357,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A +ResolutionErrors2.dfy(367,11): Warning: because of cyclic dependencies among constructor argument types, no instances of datatype 'D' can be constructed +ResolutionErrors2.dfy(364,8): Error: class 'C' with fields without known initializers, like 'a' of type 'A', must declare a constructor +ResolutionErrors2.dfy(371,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A +ResolutionErrors2.dfy(377,7): Error: cycle among redirecting types (newtypes, subset types, type synonyms): A -> B -> A +ResolutionErrors2.dfy(383,7): Error: type-synonym cycle: A -> A +ResolutionErrors2.dfy(390,25): Error: unresolved identifier: x +ResolutionErrors2.dfy(393,22): Error: unresolved identifier: x +ResolutionErrors2.dfy(396,25): Error: unresolved identifier: x +ResolutionErrors2.dfy(398,21): Error: unresolved identifier: x +ResolutionErrors2.dfy(400,21): Error: unresolved identifier: x +ResolutionErrors2.dfy(403,21): Error: unresolved identifier: x +ResolutionErrors2.dfy(410,35): Error: Wrong number of type arguments (2 instead of 1) passed to abstract type: P +ResolutionErrors2.dfy(422,13): Error: Type or type parameter is not declared in this scope: BX (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors2.dfy(432,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require bool = int) +ResolutionErrors2.dfy(437,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require B = A) +ResolutionErrors2.dfy(442,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require int = A) +ResolutionErrors2.dfy(443,6): Error: RHS (of type P) not assignable to LHS (of type P) (non-variant type parameter would require A = int) +ResolutionErrors2.dfy(448,13): Error: arguments must have comparable types (got P and P) +ResolutionErrors2.dfy(449,13): Error: arguments must have comparable types (got P and P) +ResolutionErrors2.dfy(450,13): Error: arguments must have comparable types (got P and P) +ResolutionErrors2.dfy(461,13): Error: new can be applied only to class types (got JJ) +ResolutionErrors2.dfy(469,37): Error: a set comprehension involved in a function definition is not allowed to depend on the set of allocated references, but values of 'o' (of type 'object') may contain references (see documentation for 'older' parameters) +ResolutionErrors2.dfy(470,31): Error: a set comprehension involved in a function definition is not allowed to depend on the set of allocated references, but values of 'o' (of type 'object') may contain references (see documentation for 'older' parameters) +ResolutionErrors2.dfy(477,27): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'o' +ResolutionErrors2.dfy(485,15): Error: arguments to / must be numeric or bitvector types (got set) +ResolutionErrors2.dfy(492,20): Error: a call to a possibly non-terminating method is allowed only if the calling method is also declared (with 'decreases *') to be possibly non-terminating +ResolutionErrors2.dfy(507,16): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating +70 resolution/type errors detected in ResolutionErrors2.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy new file mode 100644 index 00000000000..8ae780553da --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy @@ -0,0 +1,546 @@ +// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ------------ type variables whose values are not inferred ---------- + +module NonInferredTypeVariables { + class C { + var f: CT + } + + predicate P(x: int) + { + x < 100 + } + + ghost function Q(x: int): QT + { + var qt :| true; qt + } + + method M(n: nat) + { + var a := new MT[n]; + } + + method N(n: nat) returns (x: NT) + { + var a := new NT[10]; + x := a[3]; + } + + method DeterminedClient(n: nat) + { + ghost var q := Q(n); + var x := N(n); + var a := new array; + var c := new C; + var s: set; + var ss := new set[15]; + + q := 3.14; // this will determine the type parameter of Q to be 'real' + x := 3.14; // this will determine the type parameter of N to be 'real' + if a.Length != 0 { + a[0] := 3.14; // this will determine the type parameter of 'array' to be 'real' + } + c.f := 3.14; // this will determine the type parameter of 'C' to be 'real' + var containsPi := 3.14 in s; // this will determine the type parameter of 'set' to be 'real' + ss[12] := s; // this will determine the type parameter of 'array>' to be 'real' + } + + method BadClient(n: nat) + { + var p := P(n); // error: cannot infer the type argument for P + ghost var q := Q(n); // error: cannot infer the type argument for Q (and thus q's type cannot be determined either) + M(n); // error: cannot infer the type argument for M + var x := N(n); // error: cannot infer the type argument for N (and thus x's type cannot be determined either) + var a := new array; // error: cannot infer the type argument for 'array' + var c := new C; // error: cannot infer the type argument for 'C' + var s: set; // type argument for 'set' + var ss := new set[15]; // error: cannot infer the type argument in 'array>' + var what; // error: the type of this local variable in underspecified + } + + method MoreBadClient() + { + var b0 := forall s :: s <= {} ==> s == {}; // error: type of s underspecified + var b1 := forall s: set :: s <= {} ==> s == {}; // error: type of s underspecified + var b2 := forall c: C? :: c in {null} ==> c == null; // error: type parameter of c underspecified + + // In the following, the type of the bound variable is completely determined. + var S: set>; + ghost var d0 := forall s :: s == {7} ==> s != {}; + var d1 := forall s: set :: s in S ==> s == {}; + var ggcc0: C; + var ggcc1: C; // error: full type cannot be determined + ghost var d2 := forall c: C? :: c != null ==> c.f == 10; + ghost var d2' := forall c: C? :: c == ggcc0 && c != null ==> c.f == 10; + ghost var d2'' := forall c: C? :: c == ggcc1 && c != null ==> c.f == c.f; // error: here, type of c is not determined + + var d0' := forall s :: s == {7} ==> s != {}; + var d0'' := forall s :: s <= {7} ==> s == {}; + } +} + +// -------------- signature completion ------------------ + +module SignatureCompletion { + // datatype signatures do not allow auto-declared type parameters on the LHS + datatype Dt = Ctor(X -> Dt) // error: X is not a declared type + datatype Et = Ctor(X -> Et, Y) // error: X is not a declared type + + + method My0(s: set, x: A -> B) + method My1(x: A -> B, s: set) + method My2(s: set, x: A -> B) + method My3(x: A -> B, s: set) + + ghost function F0(s: set, x: A -> B): int + ghost function F1(x: A -> B, s: set): int + ghost function F2(s: set, x: A -> B): int + ghost function F3(x: A -> B, s: set): int +} + +// -------------- more fields as frame targets -------------------- + +module FrameTargetFields { + class C { + var x: int + var y: int + ghost var z: int + + method M() + modifies this + { + var n := 0; + ghost var save := y; + while n < x + modifies `x + { + n, x := n + 1, x - 1; + } + assert y == save; + } + + ghost method N() + modifies this + modifies `y // resolution error: cannot mention non-ghost here + modifies `z // cool + { + } + } +} + +module FrameTargetFields_More { + class C { + var x: int + var y: int + ghost var z: int + method P() + modifies this + { + ghost var h := x; + while 0 <= h + modifies `x // resolution error: cannot mention non-ghost here + modifies `z // cool + { + h, z := h - 1, 5 * z; + } + } + } +} + +// ------------------------------------------------------ + +module AmbiguousModuleReference { + module A { + module Inner { + ghost predicate Q() + } + } + + module B { + module Inner { + ghost predicate Q() + } + } + + module OpenClient { + import opened A + import opened B + lemma M() { + var a := A.Inner.Q(); // fine + var b := B.Inner.Q(); // fine + var p := Inner.Q(); // error: Inner is ambiguous (A.Inner or B.Inner) + } + } +} + +// -------------------------------------------------- + +module GhostLet { + method M() { + var x: int; + x := ghost var tmp := 5; tmp; // error: ghost -> non-ghost + x := ghost var tmp := 5; 10; // fine + x := ghost var a0, a1 :| a0 == 0 && a1 == 1; a0 + a1; // error: ghost -> non-ghost + x := ghost var a :| 0 <= a; 10; // fine + } +} + +// ------------------- tuple equality support ------------------- + +module TupleEqualitySupport { + datatype GoodRecord = GoodRecord(set<(int,int)>) + datatype BadRecord = BadRecord(set<(int, int->bool)>) // error: this tuple type does not support equality +} + +// ------------------- non-type variable names ------------------- + +module NonTypeVariableNames { + type X = int + + module Y { } + + method M(m: map) + { + assert X == X; // error (x2): type name used as variable + assert Y == Y; // error (x2): module name used as variable + assert X in m; // error (x2): type name used as variable + assert Y in m; // error (x2): module name used as variable + } + + method N(k: int) + { + assert k == X; // error (x2): type name used as variable + assert k == Y; // error (x2): module name used as variable + X := k; // error: type name used as variable + Y := k; // error: module name used as variable + } +} + +// ------------------- assign-such-that and let-such-that ------------------- + +module SuchThat { + method M() { + var x: int; + x :| 5 + 7; // error: constraint should be boolean + x :| x; // error: constraint should be boolean + var y :| 4; // error: constraint should be boolean + } + ghost function F(): int { + var w :| 6 + 8; // error: constraint should be boolean + w + } +} + +// ---------------------- NEW STUFF ---------------------------------------- + +module GhostTests { + class G { } + + method GhostNew3(n: nat) + { + var g := new G; + calc { + 5; + 2 + 3; + { if n != 0 { GhostNew3(n-1); } } // error: cannot call non-ghost method in a ghost context + 1 + 4; + { GhostNew4(g); } // error: cannot call method with nonempty modifies + -5 + 10; + } + } + + ghost method GhostNew4(g: G) + modifies g + { + } + + class MyClass { + ghost method SideEffect() + modifies this + { + } + + method NonGhostMethod() + { + } + + ghost method M() + modifies this + { + calc { + 5; + { SideEffect(); } // error: cannot call method with side effects + 5; + } + } + + ghost function F(): int + { + calc { + 6; + { assert 6 < 8; } + { NonGhostMethod(); } // error: cannot call non-ghost method + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + { MyField := 12; } // error: cannot assign to a field, and especially not a non-ghost field + { MyGhostField := 12; } // error: cannot assign to any field + { SideEffect(); } // error: cannot call (ghost) method with a modifies clause + { var x := 8; + while x != 0 + modifies this // error: cannot use a modifies clause on a loop inside a hint + { + x := x - 1; + } + } + 6; + } + 5 + } + + var MyField: int + ghost var MyGhostField: int + + method N() + { + var y := + calc { + 6; + { assert 6 < 8; } + { NonGhostMethod(); } // error: cannot call non-ghost method + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + { MyField := 12; } // error: cannot assign to a field, and especially not a non-ghost field + { MyGhostField := 12; } // error: cannot assign to any field + { M(); } // error: cannot call (ghost) method with a modifies clause + { var x := 8; + while x != 0 + modifies this // error: cannot use a modifies clause on a loop inside a hint + { + x := x - 1; + } + } + { var x := 8; + while x != 0 + { + x := x - 1; + } + } + 6; + } + 5; + } + } +} + +module CallsInStmtExpr { + class MyClass { + lemma MyLemma() + + ghost method MyEffectlessGhostMethod() + + ghost function UseLemma(): int + { + MyEffectlessGhostMethod(); // error: cannot call ghost methods (only lemmas) from this context + MyLemma(); + 10 + } + } +} + +module EvenMoreGhostTests { + ghost method NiceTry() + ensures false + { + while (true) + decreases * // error: not allowed here + { + } + } + + method BreakMayNotBeFineHere() + { + var n := 0; + var p := 0; + while (true) + { + var dontKnow; + if (n == 112) { + } else if (dontKnow == 708) { + while * { + label IfNest: + if (p == 67) { + break break; // fine, since this is not a ghost context + } else if (*) { + break break break; // error: tries to break out of more loop levels than there are + } + } + } + } + } +} + +module BadGhostTransfer { + datatype DTD_List = DTD_Nil | DTD_Cons(Car: int, Cdr: DTD_List, ghost g: int) + + method DatatypeDestructors_Ghost(d: DTD_List) { + var g1; g1 := d.g; // error: cannot use ghost member in non-ghost code + } + + method AssignSuchThatFromGhost() + { + var x: int; + ghost var g: int; + + x := g; // error: ghost cannot flow into non-ghost + + x := *; + assume x == g; // this mix of ghosts and non-ghosts is cool (but, of course, + // the compiler will complain) + + x :| x == g; // error: left-side has non-ghost, so RHS must be non-ghost as well + + x :| assume x == g; // this is cool, since it's an assume (but, of course, the + // compiler will complain) + + x :| x == 5; + g :| g <= g; + g :| assume g < g; // the compiler will complain here, despite the LHS being + // ghost -- and rightly so, since an assume is used + } +} + +module MoreGhostPrintAttempts { + method TestCalc_Ghost(m: int, n: int, a: bool, b: bool) + { + calc { + n + m; + { print n + m; } // error: non-ghost statements are not allowed in hints + m + n; + } + } +} + +module MoreLetSuchThatExpr { + method LetSuchThat_Ghost(ghost z: int, n: nat) + { + var x; x := var y :| y < z; y; // error: contraint depend on ghost (z) + } +} + +module UnderspecifiedTypedShouldBeResolvedOnlyOnce { + method CalcTest0(s: seq) { + calc { + 2; + var t :| true; 2; // error: type of 't' is underspecified + } + } +} + +module LoopResolutionTests { + class C { + var x: int + ghost var y: int + } + + ghost method M(c: C) + modifies c + { + var n := 0; + while n < 100 + modifies c`y + modifies c`x // error: not allowed to mention non-ghost field in modifies clause of ghost loops + { + c.x := c.x + 1; // error: assignment to non-ghost field not allowed here + } + } + + method MM(c: C) + modifies c + { + var n := 0; + while + invariant n <= 100 + modifies c // regression test + { + case n < 100 => n := n + 1; + } + } + + method MMX(c: C, ghost g: int) + modifies c + { + var n := 0; + while + invariant n <= 100 + modifies c`y + modifies c`x // error: not allowed to mention non-ghost field in modifies clause of ghost loops + { + case n < 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop + case g < 56 && n != 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop + } + } + + method MD0(c: C, ghost g: nat) + modifies c + decreases * + { + var n := 0; + while n + g < 100 + invariant n <= 100 + decreases * // error: disallowed on ghost loops + { + n := n + 1; // error: cannot assign to non-ghost in a ghost loop + } + } + + method MD1(c: C, ghost g: nat) + modifies c + decreases * + { + var n := 0; + while + invariant n <= 100 + decreases * // error: disallowed on ghost loops + { + case n + g < 100 => n := n + 1; // error: cannot assign to non-ghost in a ghost loop + } + } +} + +module UnderspecifiedTypesInAttributes { + function P(x: T): int + + method M() { + var {:myattr var u :| true; 6} v: int; // error: type of u is underspecified + var j {:myattr var u :| true; 6} :| 0 <= j < 100; // error: type of u is underspecified + + var a := new int[100]; + forall lp {:myattr var u :| true; 6} | 0 <= lp < 100 { // error: type of u is underspecified + a[lp] := 0; + } + + modify {:myattr P(10)} {:myattr var u :| true; 6} a; // error: type of u is underspecified + + calc {:myattr P(10)} {:myattr var u :| true; 6} // error: type of u is underspecified + { + 5; + } + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy.expect new file mode 100644 index 00000000000..c7f3e547e1b --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors3.dfy.expect @@ -0,0 +1,87 @@ +ResolutionErrors3.dfy(53,13): Error: type parameter 'PT' (inferred to be '?') in the function call to 'P' could not be determined +ResolutionErrors3.dfy(54,14): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(54,19): Error: type parameter 'QT' (inferred to be '?') in the function call to 'Q' could not be determined +ResolutionErrors3.dfy(54,20): Error: the type of this expression is underspecified +ResolutionErrors3.dfy(55,4): Error: type parameter 'MT' (inferred to be '?') to the method 'M' could not be determined +ResolutionErrors3.dfy(56,8): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(56,13): Error: type parameter 'NT' (inferred to be '?') to the method 'N' could not be determined +ResolutionErrors3.dfy(57,8): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(58,8): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(59,8): Error: the type of this local variable is underspecified +ResolutionErrors3.dfy(60,8): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(61,8): Error: the type of this local variable is underspecified +ResolutionErrors3.dfy(66,26): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(66,21): Error: type of bound variable 's' could not be determined; please specify the type explicitly +ResolutionErrors3.dfy(67,31): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(67,21): Error: type of bound variable 's' could not be determined; please specify the type explicitly +ResolutionErrors3.dfy(68,30): Error: the type of this variable is underspecified +ResolutionErrors3.dfy(68,21): Error: type of bound variable 'c' could not be determined; please specify the type explicitly +ResolutionErrors3.dfy(75,8): Error: the type of this local variable is underspecified +ResolutionErrors3.dfy(78,29): Error: type of bound variable 'c' could not be determined; please specify the type explicitly +ResolutionErrors3.dfy(89,21): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors3.dfy(90,24): Error: Type or type parameter is not declared in this scope: X (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors3.dfy(127,16): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (y) +ResolutionErrors3.dfy(144,18): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors3.dfy(174,15): Error: The name Inner ambiguously refers to a type in one of the modules A, B (try qualifying the type name with the module name) +ResolutionErrors3.dfy(184,29): Error: ghost variables such as tmp are allowed only in specification contexts. tmp was inferred to be ghost based on its declaration or initialization. +ResolutionErrors3.dfy(186,49): Error: ghost variables such as a0 are allowed only in specification contexts. a0 was inferred to be ghost based on its declaration or initialization. +ResolutionErrors3.dfy(186,54): Error: ghost variables such as a1 are allowed only in specification contexts. a1 was inferred to be ghost based on its declaration or initialization. +ResolutionErrors3.dfy(195,33): Error: set argument type must support equality (got (int, int -> bool)) +ResolutionErrors3.dfy(207,11): Error: name of type (X) is used as a variable +ResolutionErrors3.dfy(207,16): Error: name of type (X) is used as a variable +ResolutionErrors3.dfy(208,11): Error: name of module (Y) is used as a variable +ResolutionErrors3.dfy(208,16): Error: name of module (Y) is used as a variable +ResolutionErrors3.dfy(209,11): Error: name of type (X) is used as a variable +ResolutionErrors3.dfy(210,11): Error: name of module (Y) is used as a variable +ResolutionErrors3.dfy(215,16): Error: name of type (X) is used as a variable +ResolutionErrors3.dfy(216,16): Error: name of module (Y) is used as a variable +ResolutionErrors3.dfy(217,4): Error: name of type (X) is used as a variable +ResolutionErrors3.dfy(218,4): Error: name of module (Y) is used as a variable +ResolutionErrors3.dfy(228,9): Error: type of RHS of assign-such-that statement must be boolean (got int) +ResolutionErrors3.dfy(227,11): Error: type of left argument to + (int) must agree with the result type (bool) +ResolutionErrors3.dfy(227,11): Error: type of right argument to + (int) must agree with the result type (bool) +ResolutionErrors3.dfy(227,11): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) +ResolutionErrors3.dfy(229,13): Error: type of RHS of assign-such-that statement must be boolean (got int) +ResolutionErrors3.dfy(232,15): Error: type of left argument to + (int) must agree with the result type (bool) +ResolutionErrors3.dfy(232,15): Error: type of right argument to + (int) must agree with the result type (bool) +ResolutionErrors3.dfy(232,15): Error: type of + must be of a numeric type, a bitvector type, ORDINAL, char, a sequence type, or a set-like or map-like type (instead got bool) +ResolutionErrors3.dfy(248,29): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(250,17): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(275,20): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(285,24): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(298,18): Error: assignment to non-ghost field is not allowed in this context, because this is a ghost function +ResolutionErrors3.dfy(299,10): Error: a hint is not allowed to make heap updates +ResolutionErrors3.dfy(300,20): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(303,21): Error: a loop in a hint is not allowed to use 'modifies' clauses +ResolutionErrors3.dfy(322,24): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(329,18): Error: assignment to non-ghost field is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors3.dfy(330,10): Error: a hint is not allowed to make heap updates +ResolutionErrors3.dfy(331,11): Error: in a hint, calls are allowed only to lemmas +ResolutionErrors3.dfy(334,21): Error: a loop in a hint is not allowed to use 'modifies' clauses +ResolutionErrors3.dfy(360,29): Error: in a statement expression, calls are allowed only to lemmas +ResolutionErrors3.dfy(372,16): Error: a possibly infinite loop is allowed only if the enclosing method is declared (with 'decreases *') to be possibly non-terminating +ResolutionErrors3.dfy(391,12): Error: a 'break break break' statement is allowed only in contexts with 3 enclosing loops, but the current context only has 2 +ResolutionErrors3.dfy(403,20): Error: a ghost field is allowed only in specification contexts +ResolutionErrors3.dfy(411,9): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors3.dfy(417,4): Error: non-ghost variable cannot be assigned a value that depends on a ghost +ResolutionErrors3.dfy(434,8): Error: print statement is not allowed in this context (because this is a ghost method or because the statement is guarded by a specification-only expression) +ResolutionErrors3.dfy(443,29): Error: ghost variables such as z are allowed only in specification contexts. z was inferred to be ghost based on its declaration or initialization. +ResolutionErrors3.dfy(451,10): Error: the type of the bound variable 't' could not be determined +ResolutionErrors3.dfy(451,10): Error: the type of the bound variable 't' could not be determined +ResolutionErrors3.dfy(451,10): Error: the type of the bound variable 't' could not be determined +ResolutionErrors3.dfy(451,10): Error: the type of the bound variable 't' could not be determined +ResolutionErrors3.dfy(468,15): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors3.dfy(470,10): Error: assignment to non-ghost field is not allowed in this context, because this is a ghost method +ResolutionErrors3.dfy(493,15): Error: in a ghost context, only ghost fields can be mentioned as modifies frame targets (x) +ResolutionErrors3.dfy(495,25): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors3.dfy(496,35): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors3.dfy(505,4): Error: 'decreases *' is not allowed on ghost loops +ResolutionErrors3.dfy(509,8): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors3.dfy(518,4): Error: 'decreases *' is not allowed on ghost loops +ResolutionErrors3.dfy(522,29): Error: assignment to non-ghost variable is not allowed in this context, because the statement is in a ghost context; e.g., it may be guarded by a specification-only expression +ResolutionErrors3.dfy(531,21): Error: the type of the bound variable 'u' could not be determined +ResolutionErrors3.dfy(532,23): Error: the type of the bound variable 'u' could not be determined +ResolutionErrors3.dfy(535,27): Error: the type of the bound variable 'u' could not be determined +ResolutionErrors3.dfy(539,40): Error: the type of the bound variable 'u' could not be determined +ResolutionErrors3.dfy(541,38): Error: the type of the bound variable 'u' could not be determined +86 resolution/type errors detected in ResolutionErrors3.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy new file mode 100644 index 00000000000..c83a35a66ee --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy @@ -0,0 +1,332 @@ +// RUN: %exits-with 2 %verify "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ------------------- infer array types for Indexable and MultiIndexable XConstraints ---------- +// ------------------- using weaker subtyping constraints ---------- + +module AdvancedIndexableInference { + datatype MyRecord = Make(x: int, y: int) + + method M(d: array, e: seq) + requires d.Length == 100 == |e| + { + if * { + var c := d; + var xx := c[25].x; + } else if * { + var c := d; + var xx := c[25..50][10].x; + } else if * { + var c := e; + var xx := c[25].x; + } else { + var c := e; + var xx := c[25..50][10].x; + } + } +} + +// -------------------------- + +module TypeConversions { + trait J { } + class C extends J { } + + method M() returns (x: int, n: nat, o: object, j: J, c: C) { + n := x as nat; // yes, this is allowed now + o := j; + j := o; // OK for type resolution, but must be proved + j := o as J; // yes, this is allowed now + j := c; + c := j; // OK for type resolution, but must be proved + c := j as C; // yes, this is allowed now + var oo := o as realint; // error: there's no such type as "realint" (this once used to crash Dafny) + } +} + +// --------------------- regression + +module Regression_NewType { + class C { } + newtype MyInt = x: int | {} == set c: C | c // this once crashed Dafny +} + +// --------------------- regression + +module PrefixGeneratorDuplicates { + greatest predicate P() + greatest predicate P() // error: duplicate name (this once crashed Dafny) + greatest lemma L() + greatest lemma L() // error: duplicate name (this once crashed Dafny) +} + +// ------------------- unary TLA+ style predicates ------------------- + +module TLAplusOperators { + ghost function BadA(y: int): int // error: body has wrong return type + { + && 5 + y // error: using operator "&&" requires the operand to be boolean + } + + ghost function BadB(y: int): bool + { + && 5 + y // error: using operator "&&" requires the operand to be boolean + } + + ghost function BadC(y: int): int // error: body has wrong return type + { + || 5 + y // error: using operator "||" requires the operand to be boolean + } + + ghost function BadD(y: int): bool + { + || 5 + y // error: using operator "||" requires the operand to be boolean + } + + ghost function BadE(y: int): int // error: body has wrong return type + { + && (|| 5 + y) // error: bad types + } +} + +// ------------------------- divided constructors ------------------- + +module DividedConstructors { + class MyClass { + var a: nat + var b: nat + var c: nat + var n: MyClass + const t := 17 + static const g := 25 + + constructor Init(x: nat) + { + this.a := this.b; // this use of "this" in RHS is allowed + ((this)).b := 10; + n := new MyClass(); + n.a := 10; // error: not allowed use of "this" in this way + c := a + b; // error (x2): not allowed "this" in RHS + var th := this; // error: not allowed "this" in RHS + Helper(); // error: not allowed to call instance method + var mc := new MyClass(); + StaticHelper(mc); + this.StaticHelper(mc); // error: "this" is benign here; yet, it seems consistent to disallow it + StaticHelper(this); // error: cannot use "this" here + P(a); // error: cannot use "this" here + P(g); + P(this.g); // error: "this" is benign here; yet, it seems consistent to disallow it + modify this; // error: cannot use "this" here + modify this`c; // error: cannot use "this" here + modify `c; // error: cannot use (implicit) "this" here + new; + a := a + b; + Helper(); + } + + method Helper() + { + } + + static method StaticHelper(mc: MyClass) + { + } + + static method P(x: nat) + { + } + + constructor () + { + a, c := 0, 0; + new; + } + } +} + +module ConstructorsThisUsage { + class C { + var x: int + + constructor M() + requires this != null // error: cannot use "this" here + modifies this // error: cannot use "this" here (but we just issue a deprecation warning) + decreases this.x // error: cannot use "this" here + ensures this.x == 5 + { + x := 5; + } + } +} + +module ReturnBeforeNew { + class C { + var a: int + var b: int + + constructor TriesToReturnBeforeNew(xyz: int) + { + a := 0; + if xyz < 100 { + return; // error: "return" is not allowed before "new;" + } + } + } +} + +// ---------------- required auto-initialization ----------------------- + +module ZI { + // the following are different syntactic ways of saying that the type + // must support auto-initialization + type ZA(0) + type ZB(==)(0) + type ZC(0)(==) + type ZD(==,0) + type ZE(0,==) + type Y + + method P(x: G) + + method M0(a: ZA, b: ZB, c: ZC, d: ZD, e: ZE, f: F, g: G, y: Y) + { + P(a); + P(b); + P(c); + P(d); + P(e); + P(f); // error: type of argument is expected to support auto-initialization + P(g); + P(y); // error: type of argument is expected to support auto-initialization + } + + datatype List = Nil | Cons(T, List) + + method M1(xs: List, ys: List) { + P(xs); // yay, type of argument does support auto-initialization + P(ys); // yay, type of argument does support auto-initialization + } + + class Cls { + var q: int + var rs: List> + } + method M2(c: Cls?) { + P(c); + } + + newtype byte = x: int | 0 <= x < 256 // supports auto-initialization + newtype MyInt = int // supports auto-initialization + newtype SixOrMore = x | 6 <= x ghost witness 6 + newtype AnotherSixOrMore = s: SixOrMore | true ghost witness 6 + newtype MySixOrMore = x: MyInt | 6 <= x ghost witness 6 + // The resolver uses the presence/absence of a "witness" clause to figure out if the type + // supports auto-initialization. This can be inaccurate. If the type does not have a + // "witness" clause, some type replacements may slip by the resolver, but will then be + // caught by the verifier when the witness test is performed (because the witness test + // uses a zero value in the absence of a "witness" clause). + // A "ghost witness" clause tells the resolver that the type does not support + // auto-initialization, but only ghost auto-initialzation. + + newtype UnclearA = x: int | true ghost witness 0 // actually supports auto-initialization, but has a "ghost witness" clause + newtype UnclearB = x | x == 6 && x < 4 // "witness" clause omitted; type does not actually support auto-initialization + + method M3(a: byte, b: MyInt, c: SixOrMore, d: AnotherSixOrMore, e: MySixOrMore, + ua: UnclearA, ub: UnclearB) { + P(a); + P(b); + P(c); // error: type of argument is expected to support auto-initialization + P(d); // error: type of argument is expected to support auto-initialization + P(e); // error: type of argument is expected to support auto-initialization + P(ua); // error: as far as the resolver can tell, type of argument does not support auto-initialization + P(ub); // fine, as far as the resolver can tell (but this would be caught later by the verifier) + } + + type Sbyte = x: int | 0 <= x < 256 // supports auto-initialization + type SMyInt = int // supports auto-initialization + type SSixOrMore = x | 6 <= x ghost witness 6 + type SAnotherSixOrMore = s: SSixOrMore | true ghost witness 6 + type SMySixOrMore = x: SMyInt | 6 <= x ghost witness 6 + type SUnclearA = x: int | true ghost witness 0 // see note about for UnclearA + type SUnclearB = x | 6 <= x // see note about for UnclearB + + method M4(a: Sbyte, b: SMyInt, c: SSixOrMore, d: SAnotherSixOrMore, e: SMySixOrMore, + sua: SUnclearA, sub: SUnclearB) { + P(a); + P(b); + P(c); // error: type of argument is expected to support auto-initialization + P(d); // error: type of argument is expected to support auto-initialization + P(e); // error: type of argument is expected to support auto-initialization + P(sua); // error: as far as the resolver can tell, type of argument does not support auto-initialization + P(sub); // fine, as far as the resolver can tell (but this would be caught later by the verifier) + } +} + +abstract module ZI_RefinementAbstract { + type A0 + type A1 + type A2 + type A3 + type B0(0) + type B1(0) + type B2(0) + type B3(0) + type C0(00) + type C1(00) + type C2(00) + type C3(00) + + method Delta() +} + +module ZI_RefinementConcrete0 refines ZI_RefinementAbstract { + newtype Kuusi = x | 6 <= x witness 6 // supports auto-initialization + newtype Six = x | 6 <= x ghost witness 6 // does not support auto-initialization + newtype Sesis = x | 6 <= x witness * // possibly empty + type A0 = int + type A1 = Kuusi + type A2 = Six + type A3 = Sesis + type B0 = int + type B1 = Kuusi + type B2 = Six // error: RHS is expected to support auto-initialization + type B3 = Sesis // error: RHS is expected to support auto-initialization + type C0 = int + type C1 = Kuusi + type C2 = Six + type C3 = Sesis // error: RHS is expected to be nonempty +} + +module ZI_ExportSource { + export + reveals RGB + provides XYZ + datatype RGB = Red | Green | Blue + datatype XYZ = X | Y | Z +} + +module ZI_RefinementConcrete1 refines ZI_RefinementAbstract { + import Z = ZI_ExportSource + + method P(g: G) + method M(m: Z.RGB, n: Z.XYZ) { + P(m); + P(n); // error: Z.XYZ is not known to support auto-initialization + } + + type A0 + type A1(0) // error: not allowed to change auto-initialization setting + type A2(00) // error: not allowed to change nonempty setting + type B0 // error: not allowed to change auto-initialization setting + type B1(0) + type B2(00) // error: not allowed to change auto-initialization setting + type C0 // error: not allowed to change nonempty setting + type C1(0) // error: not allowed to change auto-initialization setting + type C2(00) + + method Delta< + Q, // error: not allowed to change auto-initialization setting + W, + E(0), + R(0)>() // error: not allowed to change auto-initialization setting +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy.expect new file mode 100644 index 00000000000..d23083e2e1b --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors4.dfy.expect @@ -0,0 +1,51 @@ +ResolutionErrors4.dfy(153,15): Warning: constructors no longer need 'this' to be listed in modifies clauses +ResolutionErrors4.dfy(43,19): Error: Type or type parameter is not declared in this scope: realint (did you forget to qualify a name or declare a module import 'opened'?) (note that names in outer modules are not visible in contained modules) +ResolutionErrors4.dfy(51,44): Error: range of comprehension must be of type bool (instead got C) +ResolutionErrors4.dfy(58,21): Error: Duplicate member name: P +ResolutionErrors4.dfy(60,17): Error: Duplicate member name: L +ResolutionErrors4.dfy(68,4): Error: second argument to && must be of type bool (instead got int) +ResolutionErrors4.dfy(66,17): Error: Function body type mismatch (expected int, got bool) +ResolutionErrors4.dfy(73,4): Error: second argument to && must be of type bool (instead got int) +ResolutionErrors4.dfy(78,4): Error: second argument to || must be of type bool (instead got int) +ResolutionErrors4.dfy(76,17): Error: Function body type mismatch (expected int, got bool) +ResolutionErrors4.dfy(83,4): Error: second argument to || must be of type bool (instead got int) +ResolutionErrors4.dfy(88,8): Error: second argument to || must be of type bool (instead got int) +ResolutionErrors4.dfy(86,17): Error: Function body type mismatch (expected int, got bool) +ResolutionErrors4.dfy(108,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(109,11): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(109,15): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(110,16): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(111,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(114,6): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(115,19): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(116,8): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(118,8): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(119,13): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(120,13): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(121,14): Error: in the first division of the constructor body (before 'new;'), 'this' can only be used to assign to its fields +ResolutionErrors4.dfy(152,15): Error: 'this' is not allowed in a 'static' context +ResolutionErrors4.dfy(154,16): Error: 'this' is not allowed in a 'static' context +ResolutionErrors4.dfy(171,8): Error: return statement is not allowed before 'new;' in a constructor +ResolutionErrors4.dfy(198,5): Error: type parameter (G) passed to method P must support auto-initialization (got F) (perhaps try declaring type parameter 'F' on line 191 as 'F(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors4.dfy(200,5): Error: type parameter (G) passed to method P must support auto-initialization (got Y) (perhaps try declaring abstract type 'Y' on line 187 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors4.dfy(238,5): Error: type parameter (G) passed to method P must support auto-initialization (got SixOrMore) +ResolutionErrors4.dfy(239,5): Error: type parameter (G) passed to method P must support auto-initialization (got AnotherSixOrMore) +ResolutionErrors4.dfy(240,5): Error: type parameter (G) passed to method P must support auto-initialization (got MySixOrMore) +ResolutionErrors4.dfy(241,5): Error: type parameter (G) passed to method P must support auto-initialization (got UnclearA) +ResolutionErrors4.dfy(257,17): Error: type parameter (G) passed to method P must support auto-initialization (got SSixOrMore) +ResolutionErrors4.dfy(258,24): Error: type parameter (G) passed to method P must support auto-initialization (got SAnotherSixOrMore) +ResolutionErrors4.dfy(259,19): Error: type parameter (G) passed to method P must support auto-initialization (got SMySixOrMore) +ResolutionErrors4.dfy(260,16): Error: type parameter (G) passed to method P must support auto-initialization (got SUnclearA) +ResolutionErrors4.dfy(292,7): Error: type 'B2', which does not support auto-initialization, is used to refine an abstract type that expects auto-initialization +ResolutionErrors4.dfy(293,7): Error: type 'B3', which does not support auto-initialization, is used to refine an abstract type that expects auto-initialization +ResolutionErrors4.dfy(297,7): Error: type 'C3', which may be empty, is used to refine an abstract type expected to be nonempty +ResolutionErrors4.dfy(328,4): Error: type parameter 'Q' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(331,4): Error: type parameter 'R' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(318,7): Error: type declaration 'A1' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(319,7): Error: type declaration 'A2' is not allowed to change the requirement of being nonempty +ResolutionErrors4.dfy(320,7): Error: type declaration 'B0' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(322,7): Error: type declaration 'B2' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(323,7): Error: type declaration 'C0' is not allowed to change the requirement of being nonempty +ResolutionErrors4.dfy(324,7): Error: type declaration 'C1' is not allowed to change the requirement of supporting auto-initialization +ResolutionErrors4.dfy(314,5): Error: type parameter (G) passed to method P must support auto-initialization (got Z.XYZ) +49 resolution/type errors detected in ResolutionErrors4.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy new file mode 100644 index 00000000000..5a1973feb51 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy @@ -0,0 +1,640 @@ +// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ----- constructor-less classes with need for initialization ----- + +module ConstructorlessClasses { + class C { // error: must have constructor + var x: int + var s: string + var t: set + var u: T + const c: T + } + + method Test() + { + var c := new C>; + print "int: ", c.x, "\n"; + print "string: ", c.s, "\n"; + print "set>: ", c.t, "\n"; + print "real: ", c.u, "\n"; + print "real: ", c.c, "\n"; + } + + codatatype Co = Suc(Co>) // does not know a known compilable value + codatatype Co2 = CoEnd | CoSuc(Co2) + + trait Trait { + var co: Co // has no known initializer + } + + class Class extends Trait { // error: must have constructor, because of inherited field "co" + } + + class CoClass0 { // error: must have constructor + const co: Co + } + + class CoClass1 { // fine + const co: Co2 := CoEnd + } + + trait CoTrait { + const co: Co2 := CoEnd + } + + class CoClass2 extends CoTrait { // fine + } + + iterator Iter(u: T) yields (v: T) + { + } +} + +module GhostWitness { + type BadGhost_EffectlessArrow = f: A -> B + | true + witness (GhostEffectlessArrowWitness) // error: a ghost witness must use the keyword "ghost" + + type GoodGhost_EffectlessArrow = f: A -> B + | true + ghost witness (GhostEffectlessArrowWitness) + + ghost function GhostEffectlessArrowWitness(a: A): B + { + var b: B :| true; b + } +} + +module BigOrdinalRestrictions { // also see BigOrdinalRestrictionsExtremePred below + method Test() { + var st: set; // error: cannot use ORDINAL as type argument + var p: (int, ORDINAL); // error: cannot use ORDINAL as type argument + var o: ORDINAL; // okay + ghost var f := F(o); // error: cannot use ORDINAL as type argument + f := F'(); // error: cannot use ORDINAL as type argument + f := F'<(char,ORDINAL)>(); // error: cannot use ORDINAL as type argument + var lambda := F'; // error: cannot use ORDINAL as type argument + ParameterizedMethod(o); // error: cannot use ORDINAL as type argument + assert forall r: ORDINAL :: P(r); + assert forall r: (ORDINAL, int) :: F(r.1) < 8; + assert exists x: int, r: ORDINAL, y: char :: P(r) && F(x) == F(y); + var s := set r: ORDINAL | r in {}; // error: cannot use ORDINAL as type argument (to set) + var s' := set r: ORDINAL | true :: 'G'; + ghost var m := imap r: ORDINAL :: 10; + var sq := [o, o]; // error: cannot use ORDINAL as type argument (to seq) + var mp0 := map[o := 'G']; // error: cannot use ORDINAL as type argument (to map) + var mp1 := map['G' := o]; // error: cannot use ORDINAL as type argument (to map) + var w := var h: ORDINAL := 100; h + 40; // okay + var w': (int, ORDINAL); // error: cannot use ORDINAL as type argument + var u: ORDINAL :| u == 15; + var ti: ORDINAL :| assume true; + var u': (ORDINAL, int) :| u' == (15, 15); // error (x2): ORDINAL cannot be a type argument + var ti': (ORDINAL, ORDINAL) :| assume true; // error (x4): ORDINAL cannot be a type argument + var lstLocal := var lst: ORDINAL :| lst == 15; lst; + var lstLocal' := var lst: (ORDINAL, int) :| lst == (15, 15); lst.1; // error: ORDINAL cannot be a type argument + + if yt: ORDINAL :| yt == 16 { + ghost var pg := P(yt); + } + + if { + case zt: ORDINAL :| zt == 180 => + ghost var pg := P(zt); + } + + forall om: ORDINAL // allowed + ensures om < om+1 + { + } + + var arr := new int[23]; + forall om: ORDINAL | om == 11 // allowed + { + arr[0] := 0; + } + } + + ghost function F(g: G): int + ghost function F'(): int + method ParameterizedMethod(g: G) + ghost predicate P(g: ORDINAL) +} + +module IteratorDuplicateParameterNames { + // each of the following once caused a crash in the resolver + iterator MyIterX(u: char) yields (u: char) // error: duplicate name "u" + iterator MyIterY(us: char) yields (u: char) // error: in-effect-duplicate name "us" +} + +module TernaryTypeCheckinngAndInference { + codatatype Stream = Cons(int, Stream) + + method M(k: nat, K: ORDINAL, A: Stream, B: Stream) + requires A == B + { + // all of the following are fine + assert A ==#[k] B; + assert A ==#[K] B; + assert A ==#[3] B; + var b; + assert A ==#[b] B; + } +} + +module DontQualifyWithNonNullTypeWhenYouMeanAClass { + module Z { + class MyClass { + static const g := 100 + } + + method M0() { + var x := MyClass?.g; // error: use MyClass, not MyClass? + assert x == 100; + } + + method M1() { + var x := MyClass.g; // that's it! + assert x == 100; + } + + method P(from: MyClass) returns (to: MyClass?) { + to := from; + } + + method Q() { + var x := MyClass; // error: type used as variable + var y := MyClass?; // error: type used as variable + } + } + + module A { + class MyClass { + static const g := 100 + } + } + + module B { + import A + + method M0() { + var x := A.MyClass?.g; // error: use MyClass, not MyClass? + assert x == 100; + } + + method M1() { + var x := A.MyClass.g; // that's it! + assert x == 100; + } + + method P(from: A.MyClass) returns (to: A.MyClass?) { + to := from; + } + + method Q() { + var x := A.MyClass; // error: type used as variable + var y := A.MyClass?; // error: type used as variable + } + } +} + +module UninterpretedModuleLevelConst { + type Six = x | 6 <= x witness 6 + type Odd = x | x % 2 == 1 ghost witness 7 + const S: Six // fine + const X: Odd // error: the type of a non-ghost static const must have a known initializer + + class MyClass { } + const Y: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer + ghost const Y': MyClass // error: the type of a ghost static const must be known to be nonempty + + class AnotherClass { // fine, the class itself is not required to have a constructor, because the bad fields are static + static const k := 18 + static const W: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer + static const U: Six // fine, since Six has a non-ghost witness and thus has a known initializer + static const O: Odd // error: the type of a non-ghost static const must have a known initializer + const u: Six + } + + trait Trait { + static const k := 18 + static const W: MyClass // error: the type of a non-ghost static const must have a known (non-ghost) initializer + static const U: Six // fine, since Six has a non-ghost witness and thus has a known initializer + static const O: Odd // error: the type of a non-ghost static const must have a known initializer + } + + class ClassyTrait extends Trait { // fine, since the bad fields in Trait are static + } + + trait InstanceConst { + const w: MyClass + } + + class Instance extends InstanceConst { // error: because of "w", must declare a constructor + } + + trait GhostTr { + ghost const w: MyClass // the responsibility to initialize "w" lies with any class that implements "GhostTr" + } + + class GhostCl extends GhostTr { // error: w and z need initialization, so GhostCl must have a constructor + ghost const z: MyClass + } +} + +module NonThisConstAssignments { + const X: int + + class Cla { + constructor () { + X := 15; // error: can never assign to static const (this used to crash) + Clb.Y := 15; // error: can never assign to static const (this used to crash) + } + } + + class Clb { + static const Y: int + + constructor () { + Y := 15; // error: cannot ever assign to a static const (this used to crash) + } + } + + class Clc { + const Z: int + + constructor (c: Clc) { + c.Z := 15; // error: can assign to const only for 'this' (this used to crash) + } + } + + class Cld { + const Z: int + + constructor () { + Z := 15; + } + } +} + +module ConstGhostRhs { + class S { + const m: int := n // error: use of ghost to assign non-ghost field + ghost const n: int + } + + const a: int := b // error: use of ghost to assign non-ghost field + ghost const b: int + + class S' { + ghost const m': int := n' + const n': int + } + + ghost const a': int := b' + const b': int + +} + +module Regression15 { + predicate F(i: int, j: int) { true } + function S(i: int): set { {i} } + + method M0() returns (b: bool) { + b := forall i, j | j <= i <= 100 && i <= j < 100 :: true; // error: this bogus cyclic dependency was once allowed + } + + method M4() returns (b: bool) { + b := forall i, j :: j <= i < 100 && j in S(i) ==> F(i,j); // error: this bogus cyclic dependency was once allowed + } +} + +module AllocDepend0a { + class Class { + const z := if {} == set c: Class | true then 5 else 4 // error: condition depends on alloc (also not compilable, but that's not reported in the same pass) + } + + const y := if {} == set c: Class | true then 5 else 4 // error: condition depends on alloc (also not compilable, but that's not reported in the same pass) + + newtype byte = x | x < 5 || {} == set c: Class | true // error: condition not allowed to depend on alloc + type small = x | x < 5 || {} == set c: Class | true // error: condition not allowed to depend on alloc +} + +module AllocDepend0b { + class Class { } + + method M() returns (y: int, z: int) { + z := if {} == set c: Class | true then 5 else 4; // error: not compilable + y := if {} == set c: Class | true then 5 else 4; // error: not compilable + } +} + +module AllocDepend1 { + class Class { } + + predicate P(x: int) { + x < 5 || {} == set c: Class | true // error: function not allowed to depend on alloc + } +} + +module AllocDepend2 { + class Klass { + const z := if exists k: Klass :: allocated(k) then 3 else 4 // error (x2): condition depends on alloc + } + + const y := if exists k: Klass :: allocated(k) then 3 else 4 // error (x2): condition depends on alloc + + newtype byte = x | x < 5 || exists k: Klass :: allocated(k) // error: condition not allowed to depend on alloc + type small = x | x < 5 || exists k: Klass :: allocated(k) // error: condition not allowed to depend on alloc +} + +module AllocDepend3 { + class Klass { } + + predicate P(x: int) { + x < 5 || exists k: Klass :: allocated(k) // error: function not allowed to depend on alloc + } +} + +module AllocDepend4 { + class Xlass { + const z := if var k: Xlass? := null; allocated(k) then 3 else 4 // error (x2): condition depends on alloc + } + + const y := if var k: Xlass? := null; allocated(k) then 3 else 4 // error (x2): condition depends on alloc + + newtype byte = x | x < 5 || var k: Xlass? := null; allocated(k) // error: condition not allowed to depend on alloc + type small = x | x < 5 || var k: Xlass? := null; allocated(k) // error: condition not allowed to depend on alloc +} + +module AllocDepend5 { + class Xlass { } + + predicate P(x: int) { + x < 5 || var k: Xlass? := null; allocated(k) // error: function not allowed to depend on alloc + } +} + +module AbstemiousCompliance { + codatatype EnormousTree = Node(left: EnormousTree, val: X, right: EnormousTree) + + ghost function {:abstemious} Five(): int { + 5 // error: an abstemious function must return with a co-constructor + } + + ghost function {:abstemious} Id(t: EnormousTree): EnormousTree + { + t // to be abstemious, a parameter is as good as a co-constructor + } + + ghost function {:abstemious} IdGood(t: EnormousTree): EnormousTree + { + match t + case Node(l, x, r) => Node(l, x, r) + } + + ghost function {:abstemious} AlsoGood(t: EnormousTree): EnormousTree + { + Node(t.left, t.val, t.right) + } + + ghost function {:abstemious} UniformTree(x: X): EnormousTree + { + Node(UniformTree(x), x, UniformTree(x)) + } + + ghost function {:abstemious} AlternatingTree(x: X, y: X): EnormousTree + { + Node(AlternatingTree(y, x), x, AlternatingTree(y, x)) + } + + ghost function {:abstemious} AnotherAlternatingTree(x: X, y: X): EnormousTree + { + var t := Node(AlternatingTree(x, y), y, AlternatingTree(x, y)); + Node(t, x, t) + } + + ghost function {:abstemious} NonObvious(x: X): EnormousTree + { + AlternatingTree(x, x) // error: does not return with a co-constructor + } + + ghost function {:abstemious} BadDestruct(t: EnormousTree): EnormousTree + { + Node(t.left, t.val, t.right.right) // error: cannot destruct t.right + } + + ghost function {:abstemious} BadMatch(t: EnormousTree): EnormousTree + { + match t.right // error: cannot destruct t.right + case Node(a, x, b) => + Node(a, x, b) + } + + ghost function {:abstemious} BadEquality(t: EnormousTree, u: EnormousTree, v: EnormousTree): EnormousTree + { + if t == u then // error: cannot co-compare + Node(t.left, t.val, t.right) + else if u != v then // error: cannot co-compare + Node(u.left, u.val, u.right) + else + Node(v.left, v.val, v.right) + } + + ghost function {:abstemious} Select(b: bool, t: EnormousTree, u: EnormousTree): EnormousTree + { + if b then t else u // abstemious yes: parameters are as good as a co-constructors + } + + ghost function {:abstemious} Select'(b: bool, t: EnormousTree, u: EnormousTree): EnormousTree + { + if b then + Node(t.left, t.val, t.right) // fine, too + else + Node(u.left, u.val, u.right) // fine, too + } +} + +module BigOrdinalRestrictionsExtremePred { + least predicate Test() { + var st: set := {}; // error: cannot use ORDINAL as type argument + var p: (int, ORDINAL) := (0,0); // error: cannot use ORDINAL as type argument + var o: ORDINAL := 0; // okay + ghost var f := F(o); // error: cannot use ORDINAL as type argument + var f := F'(); // error: cannot use ORDINAL as type argument + var f := F'<(char,ORDINAL)>(); // error: cannot use ORDINAL as type argument + var lambda := F'; // error: cannot use ORDINAL as type argument + ParameterizedLemma(o); // error: cannot use ORDINAL as type argument + assert forall r: ORDINAL :: P(r); // error: cannot quantify over ORDINAL here + assert forall r: (ORDINAL, int) :: F(r.1) < 8; // error: cannot quantify over ORDINAL here + assert exists x: int, r: ORDINAL, y: char :: P(r) && F(x) == F(y); // error: cannot quantify over ORDINAL here + var s := set r: ORDINAL | r in {}; // error (x2): cannot use ORDINAL as type argument (to set) + var s' := set r: ORDINAL | true :: 'G'; // error: cannot use ORDINAL here + ghost var m := imap r: ORDINAL :: 10; // error (x2): cannot use ORDINAL here + var sq := [o, o]; // error: cannot use ORDINAL as type argument (to seq) + var mp0 := map[o := 'G']; // error: cannot use ORDINAL as type argument (to map) + var mp1 := map['G' := o]; // error: cannot use ORDINAL as type argument (to map) + var w := var h: ORDINAL := 100; h + 40; // okay + var w': (int, ORDINAL) := (8,8); // error: cannot use ORDINAL as type argument + var u: ORDINAL :| u == 15; + var ti: ORDINAL :| true; + var u': (ORDINAL, int) :| u' == (15, 15); // error (x2): ORDINAL cannot be a type argument + var ti': (ORDINAL, ORDINAL) :| true; // error (x2): ORDINAL cannot be a type argument + var lstLocal := var lst: ORDINAL :| lst == 15; lst; + var lstLocal' := var lst: (ORDINAL, int) :| lst == (15, 15); lst.1; // error: ORDINAL cannot be a type argument + + var gr := if yt: ORDINAL :| yt == 16 then + ghost var pg := P(yt); 5 + else + 7; + + calc { + 100; + == { + forall om: ORDINAL // allowed + ensures om < om+1 + { + } + } + 100; + } + + true + } + + ghost function F(g: G): int + ghost function F'(): int + lemma ParameterizedLemma(g: G) + ghost predicate P(g: ORDINAL) +} + +// ----- label domination ----- + +module LabelDomination { + method DuplicateLabels(n: int) { + var x; + if (n < 7) { + label L: x := x + 1; + } else { + label L: x := x + 1; + } + assert old@L(true); // error: L is not available here + label L: x := x + 1; + assert old@L(true); + label L: x := x + 1; // error: duplicate label + assert old@L(true); + + { + label K: + x := x + 1; + } + assert old@K(true); + label K: // error: duplicate label + assert old@K(true); + } + + datatype Color = A | B | C + method Branches(n: int, c: Color) { + var x: int; + if n < 2 { + label X: x := x + 1; + } else if n < 4 { + label X: x := x + 1; + } else { + label X: x := x + 1; + } + if * { + label X: x := x + 1; + } else { + label X: x := x + 1; + } + + if { + case true => + label X: x := x + 1; + case true => + label X: x := x + 1; + } + + var i := 0; + while i < x { + label X: x := x + 1; + i := i + 1; + } + + i := 0; + while { + case i < x => + label X: x := x + 1; + i := i + 1; + case i < x => + label X: x := x + 1; + i := i + 1; + } + + match c { + case A => + label X: x := x + 1; + case B => + label X: x := x + 1; + case C => + label X: x := x + 1; + } + + label X: x := x + 1; // all okay + } + + method A0() { + label Q: + assert true; + } + + method A1() { + label Q: + assert true; + } + + class MyClass { + var x: int + + method LabelNotInScope_Old(y: int) { + label GoodLabel: + if y < 5 { + label Treasure: + assert true; + } + assert old(x) == old@Treasure(x); // error: no label Treasure in scope + assert 10 == old@WonderfulLabel(x); // error: no label WonderfulLabel in scope + assert old(x) == old@GoodLabel(x); + assert old(x) == old@FutureLabel(x); // error: no label FutureLabel in scope + label FutureLabel: {} + } + + method LabelNotInScope_Unchanged(y: int) { + label GoodLabel: + if y < 5 { + label Treasure: + assert true; + } + assert unchanged@Treasure(`x); // error: no label Treasure in scope + assert unchanged@WonderfulLabel(this); // error: no label WonderfulLabel in scope + assert unchanged@GoodLabel(this); + assert unchanged@FutureLabel(this); // error: no label FutureLabel in scope + label FutureLabel: {} + } + + method LabelNotInScope_Fresh(y: int, c: MyClass) { + label GoodLabel: + if y < 5 { + label Treasure: + assert true; + } + assert fresh@Treasure(c); // error: no label Treasure in scope + assert fresh@WonderfulLabel(c); // error: no label WonderfulLabel in scope + assert fresh@GoodLabel(c); + assert fresh@FutureLabel(c); // error: no label FutureLabel in scope + label FutureLabel: {} + } + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy.expect new file mode 100644 index 00000000000..1a41d63271f --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors5.dfy.expect @@ -0,0 +1,107 @@ +ResolutionErrors5.dfy(50,33): Error: type of yield-parameter must support auto-initialization (got 'T') +ResolutionErrors5.dfy(58,13): Error: a ghost function is allowed only in specification contexts +ResolutionErrors5.dfy(72,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(73,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(75,19): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F' is not allowed to use ORDINAL +ResolutionErrors5.dfy(76,9): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(77,9): Error: type parameter 'G' (passed in as '(char, ORDINAL)') to function call 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(78,18): Error: type parameter 'G' (passed in as 'ORDINAL') to the function 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(79,4): Error: type parameter 'G' (passed in as 'ORDINAL') to the method 'ParameterizedMethod' is not allowed to use ORDINAL +ResolutionErrors5.dfy(83,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(85,14): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(86,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(87,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(88,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(90,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(93,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(93,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(94,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(94,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(94,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(94,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(96,25): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(127,36): Error: Name of yield-parameter is used by another member of the iterator: u +ResolutionErrors5.dfy(128,37): Error: Name of implicit yield-history variable 'u' is already used by another member of the iterator +ResolutionErrors5.dfy(153,15): Error: To access members of class 'MyClass', write 'MyClass', not 'MyClass?' +ResolutionErrors5.dfy(167,15): Error: name of type (MyClass) is used as a variable +ResolutionErrors5.dfy(168,15): Error: name of type (MyClass?) is used as a variable +ResolutionErrors5.dfy(182,17): Error: To access members of class 'MyClass', write 'MyClass', not 'MyClass?' +ResolutionErrors5.dfy(196,17): Error: name of type (MyClass) is used as a variable +ResolutionErrors5.dfy(197,17): Error: name of type (MyClass?) is used as a variable +ResolutionErrors5.dfy(206,8): Error: static non-ghost const field 'X' of type 'Odd' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(209,8): Error: static non-ghost const field 'Y' of type 'MyClass' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(210,14): Error: static ghost const field 'Y'' of type 'MyClass' (which may be empty) must give a defining value +ResolutionErrors5.dfy(214,17): Error: static non-ghost const field 'W' of type 'MyClass' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(216,17): Error: static non-ghost const field 'O' of type 'Odd' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(222,17): Error: static non-ghost const field 'W' of type 'MyClass' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(224,17): Error: static non-ghost const field 'O' of type 'Odd' (which does not have a default compiled value) must give a defining value +ResolutionErrors5.dfy(234,8): Error: class 'Instance' with fields without known initializers, like 'w' of type 'MyClass', must declare a constructor +ResolutionErrors5.dfy(241,8): Error: class 'GhostCl' with fields without known initializers, like 'z' of type 'MyClass', must declare a constructor +ResolutionErrors5.dfy(251,6): Error: LHS of assignment must denote a mutable field +ResolutionErrors5.dfy(252,10): Error: LHS of assignment must denote a mutable field +ResolutionErrors5.dfy(260,6): Error: LHS of assignment must denote a mutable field +ResolutionErrors5.dfy(268,8): Error: LHS of assignment must denote a mutable field of 'this' +ResolutionErrors5.dfy(287,18): Error: a ghost const field is allowed only in specification contexts +ResolutionErrors5.dfy(283,20): Error: a ghost const field is allowed only in specification contexts +ResolutionErrors5.dfy(305,9): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' +ResolutionErrors5.dfy(309,9): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' +ResolutionErrors5.dfy(318,22): Error: a set comprehension involved in a const field definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) +ResolutionErrors5.dfy(315,24): Error: a set comprehension involved in a const field definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) +ResolutionErrors5.dfy(320,36): Error: a set comprehension involved in a newtype definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) +ResolutionErrors5.dfy(321,34): Error: a set comprehension involved in a subset type definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) +ResolutionErrors5.dfy(328,18): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'c' +ResolutionErrors5.dfy(329,18): Error: set comprehensions in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'c' +ResolutionErrors5.dfy(337,19): Error: a set comprehension involved in a predicate definition is not allowed to depend on the set of allocated references, but values of 'c' (of type 'Class') may contain references (see documentation for 'older' parameters) +ResolutionErrors5.dfy(346,35): Error: a const field definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(343,37): Error: a const field definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(348,49): Error: a newtype definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(349,47): Error: a subset type definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(356,32): Error: a predicate definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(365,39): Error: a const field definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(362,41): Error: a const field definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(367,53): Error: a newtype definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(368,51): Error: a subset type definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(375,36): Error: a predicate definition is not allowed to depend on the set of allocated references +ResolutionErrors5.dfy(382,31): Error: the value returned by an abstemious function must come from invoking a co-constructor +ResolutionErrors5.dfy(418,31): Error: the value returned by an abstemious function must come from invoking a co-constructor +ResolutionErrors5.dfy(425,32): Error: an abstemious function is allowed to invoke a codatatype destructor only on its parameters +ResolutionErrors5.dfy(430,12): Error: an abstemious function is allowed to codatatype-match only on its parameters +ResolutionErrors5.dfy(437,9): Error: an abstemious function is not allowed to check codatatype equality +ResolutionErrors5.dfy(439,14): Error: an abstemious function is not allowed to check codatatype equality +ResolutionErrors5.dfy(464,19): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F' is not allowed to use ORDINAL +ResolutionErrors5.dfy(465,13): Error: type parameter 'G' (passed in as 'ORDINAL') to function call 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(466,13): Error: type parameter 'G' (passed in as '(char, ORDINAL)') to function call 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(467,18): Error: type parameter 'G' (passed in as 'ORDINAL') to the function 'F'' is not allowed to use ORDINAL +ResolutionErrors5.dfy(468,4): Error: type parameter 'G' (passed in as 'ORDINAL') to the lemma 'ParameterizedLemma' is not allowed to use ORDINAL +ResolutionErrors5.dfy(469,18): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(470,18): Error: type of bound variable 'r' ('(ORDINAL, int)') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(471,26): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(472,17): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(473,18): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(474,24): Error: type of bound variable 'r' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(485,25): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(487,17): Error: type of bound variable 'yt' ('ORDINAL') is not allowed to use type ORDINAL +ResolutionErrors5.dfy(483,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(483,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(482,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(479,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(477,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(476,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(475,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(474,14): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(472,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(462,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(461,8): Error: an ORDINAL type is not allowed to be used as a type argument +ResolutionErrors5.dfy(522,11): Error: no label 'L' in scope at this time +ResolutionErrors5.dfy(525,10): Error: label shadows a dominating label +ResolutionErrors5.dfy(533,10): Error: label shadows a dominating label +ResolutionErrors5.dfy(607,23): Error: no label 'Treasure' in scope at this time +ResolutionErrors5.dfy(608,19): Error: no label 'WonderfulLabel' in scope at this time +ResolutionErrors5.dfy(610,23): Error: no label 'FutureLabel' in scope at this time +ResolutionErrors5.dfy(620,13): Error: no label 'Treasure' in scope at this time +ResolutionErrors5.dfy(621,13): Error: no label 'WonderfulLabel' in scope at this time +ResolutionErrors5.dfy(623,13): Error: no label 'FutureLabel' in scope at this time +ResolutionErrors5.dfy(633,13): Error: no label 'Treasure' in scope at this time +ResolutionErrors5.dfy(634,13): Error: no label 'WonderfulLabel' in scope at this time +ResolutionErrors5.dfy(636,13): Error: no label 'FutureLabel' in scope at this time +106 resolution/type errors detected in ResolutionErrors5.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy new file mode 100644 index 00000000000..e4970a3baf0 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy @@ -0,0 +1,388 @@ +// RUN: %exits-with 2 %verify "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// ----- bad use of types without auto-initializers ----- + +module Initialization { + datatype Yt = MakeYt(x: int, y: Y) + type Even = x | x % 2 == 0 + type Odd = x | x % 2 == 1 witness 17 + type GW = x | x % 2 == 1 ghost witness 17 + + method DefiniteAssignmentViolation() returns (e: Yt, o: Yt, g: Yt) + { + } // no resolution errors (but verification errors, see NonZeroInitialization.dfy) + + method ArrayElementInitViolation() returns (e: array>, o: array>, g: array>) + { + e := new Yt[20]; + o := new Yt[20]; + g := new Yt[20]; + } // no resolution errors (but verification errors, see NonZeroInitialization.dfy) + + method GimmieOne() returns (g: G) + { + } + + method TypeParamViolation() returns (e: Yt, o: Yt, g: Yt) + { + e := GimmieOne>(); + o := GimmieOne>(); + g := GimmieOne>(); // error: cannot pass Yt to a (0)-parameter + } +} + +// ----------------- regression tests ---------------------------------------------------- + +module FreshTypeInferenceRegression { + class MyClass { + method M(N: nat) + { + var i, os := 0, {}; + while i < N + invariant fresh(os) + invariant forall o :: o in os ==> fresh(o.inner) // error: type of "o" not yet known (this once caused a crash) + { + var o := new Outer(); + os, i := os + {o}, i + 1; + } + } + } + + class Outer { + const inner: Inner + + constructor () + ensures fresh(inner) + { + inner := new Inner(); + } + } + + class Inner { + constructor () + } +} + +module RegressionTest { + class Cache { + method Lookup(K: X) returns (V: X) + { + V := Cache[K]; // error: Cache is not a field but a type + } + } +} + +module ExistsImpliesWarning { + method M(a: array, b: array) + requires a.Length == b.Length + requires exists i :: 0 <= i < a.Length ==> a[i] == b[i] // warning + requires exists i :: true && (0 <= i < a.Length ==> a[i] == b[i]) + requires exists i :: (0 <= i < a.Length ==> a[i] == b[i]) + requires exists i | 0 <= i < a.Length :: true ==> a[i] == b[i] + requires exists i | 0 <= i < a.Length :: a[i] == b[i] + requires exists i | 0 <= i < a.Length ==> a[i] == b[i] :: true + requires exists i :: !(0 <= i < a.Length) || a[i] == b[i] + requires exists i :: a[i] == b[i] <== 0 <= i < a.Length // warning + requires exists i :: !(0 <= i < a.Length && a[i] != b[i]) + requires exists i :: 0 <= i < a.Length && a[i] == b[i] + requires exists i :: true ==> (0 <= i < a.Length && a[i] == b[i]) // warning + { + } + + method N(a: array, b: array) + requires a.Length == b.Length + requires forall i :: 0 <= i < a.Length ==> a[i] == b[i] + requires forall i :: true && (0 <= i < a.Length ==> a[i] == b[i]) + requires forall i :: (0 <= i < a.Length ==> a[i] == b[i]) + requires forall i | 0 <= i < a.Length :: true ==> a[i] == b[i] + requires forall i | 0 <= i < a.Length :: a[i] == b[i] + requires forall i | 0 <= i < a.Length ==> a[i] == b[i] :: true + requires forall i :: !(0 <= i < a.Length) || a[i] == b[i] + requires forall i :: a[i] == b[i] <== 0 <= i < a.Length + requires forall i :: !(0 <= i < a.Length && a[i] != b[i]) + requires forall i :: 0 <= i < a.Length && a[i] == b[i] + requires forall i :: true ==> (0 <= i < a.Length && a[i] == b[i]) + { + } +} + +// --------------- ghost (regression) tests, receivers ------------------------------------- + +module GhostReceiverTests { + class C { + ghost function F(x: int): int { 3 } + function G(x: int): int { 4 } + lemma L(x: int) { } + method M(x: int) { } + } + + method Caller(x: int, ghost z: int, c: C, ghost g: C) { + { + var y; + y := c.F(x); // error: LHS is non-ghost, so RHS cannot use ghost function F + y := g.F(x); // error: LHS is non-ghost, so RHS cannot use ghost function F + y := c.G(x); + y := g.G(x); // error: LHS is non-ghost, so RHS cannot use ghost variable g + } + { + // all of the these are fine, because: the LHS is ghost and, therefore, the whole statement is + ghost var y; + y := c.F(x); + y := g.F(x); + y := c.G(x); + y := g.G(x); + } + { + // all of the these are fine, because: the LHS is ghost and, therefore, the whole statement is + ghost var y; + y := c.F(z); + y := g.F(z); + y := c.G(z); + y := g.G(z); + } + c.L(x); + g.L(x); + c.M(x); + g.M(x); // error: cannot pass ghost receiver to compiled method + } +} + +// --------------- ghost RHS of constants (regression) tests ------------------------------------- + +module GhostRhsConst { + class C { + ghost function F(n: nat): nat { n } // a ghost function + static ghost function G(n: nat): nat { n } // a ghost function + const b := F(0) // error: RHS uses a ghost function + static const u := G(0) // error: RHS uses a ghost function + } + + trait R { + ghost function F(n: nat): nat { n } // a ghost function + static ghost function G(n: nat): nat { n } // a ghost function + const b := F(0) // error: RHS uses a ghost function + static const u := G(0) // error: RHS uses a ghost function + } +} + +// --------------- errors from nested modules ------------------------------------- + +module ErrorsFromNestedModules { + method M() { + U.V.Test(); + UU.V.Test(); + } + + module U { // regression test: since U is rather empty, this had once caused an error + module V { + method Test() { + } + module W { + const x1 := 12 * false // error: bad types + } + } + } + + module UU.V { // same regression as above + method Test() { + } + module W { + const x1 := 12 * false // error: bad types + } + } +} + +// --------------- name clashes related to prefix-named modules ------------------------------------- + +module NameClashes { + module U.G { + } + module U { + class G { } // error: duplicate name: G + class H { } + } + module U.H { // error: duplicate name: H + } +} + +// --------------- regression ghost tests ------------------------------------- + +module RegressionGhostTests { + class Cell { + var data: int + } + + method field(x: Cell) + modifies x + { + ghost var y := x; + x.data := 42; + y.data := 42; // error: assignment to non-ghost field depends on a ghost + (assert x == y; x).data := 42; + (assert x == y; y).data := 42; // error: assignment to non-ghost field depends on a ghost + } + + method arr(a: array) + requires 5 < a.Length + modifies a + { + ghost var b := a; + ghost var i := 5; + a[i] := 42; // error: assignment to non-ghost field depends on a ghost + b[5] := 42; // error: assignment to non-ghost field depends on a ghost + } + + method arr2(a: array2) + requires 5 < a.Length0 && 5 < a.Length1 + modifies a + { + ghost var b := a; + ghost var i := 5; + a[i,5] := 42; // error: assignment to non-ghost field depends on a ghost + a[5,i] := 42; // error: assignment to non-ghost field depends on a ghost + b[5,5] := 42; // error: assignment to non-ghost field depends on a ghost + } +} + +// --------------- regression test const in frame expression ------------------------------ + +module RegressionConstFrameExpression { + class C { + const x: int + var y: int + } + method m(c: C) + modifies c`x + modifies c`y + ensures unchanged(c`x) + ensures unchanged(c) + { + + } +} + +// --------------- change in language semantics to forbid !! on maps ------------------------------ + +module MapDisjointnessNoMore { + method M(a: map, b: map) { + assert a !! b; // error: !! is (no longer) support on maps + assert a.Keys !! b.Keys; // instead, this is the way to do it + } +} + +// --------------- expect statements ------------------------------ + +module ExpectStatements { + ghost function UnsafeDivide(a: int, b: int): int { + expect b != 0; // expect statement is not allowed in this context + a / b + } + + method M() { + ghost var g := 5; + expect forall i : int :: i == i; // error: quantifiers in non-ghost contexts must be compilable + expect false, if g == 5 then "boom" else "splat"; // error: ghost variables are allowed only in specification contexts + } +} + +// --------------- type-parameter scopes ------------------------------ + +module TypeParameterScopes { + class C { + function G(): X + + method M(f: X) { + var h: X := f; + var k: X := G(); // error: this is the wrong X + } + + function F(f: X): int { + var h: X := f; + var k: X := G(); // error: this is the wrong X + 10 + } + } +} + +// --------------- type of function members (regression tests) ------------------------------ + +module TypeOfFunctionMember { + ghost function Fo(x: X): int + + lemma M() { + // Both of the following once crashed the type checker + var rd := Fo.reads; + var rq := Fo.requires; + } +} + +// --------------- update operations ------------------------------ + +module CollectionUpdates { + // Update operations on collections must have the right types, modulo subset types. + // For verification errors, see Maps.dfy. + trait Trait { } + class Elem extends Trait { } + + method UpdateValiditySeq(d: Trait, e: Elem) { + var s: seq := [e, e, e, e, e]; + s := s[1 := d]; // error: d is not an Elem (and is not a subset type of it, either) + } + + method UpdateValidityMultiset(d: Trait) { + var s: multiset; + s := s[d := 5]; // error: element value is not a Elem + } + + method UpdateValidityMap(d: Trait, e: Elem) { + var m: map; + if * { + m := m[d := e]; // error: key is not a Elem + } else { + m := m[e := d]; // error: value is not a Elem + } + } +} + +// --------------- update operations ------------------------------ + +module MoreAutoInitAndNonempty { + type A(0) + type B(00) + type C + + method Q(f: F) + method P(g: G) + method R(h: H) + + function FQ(f: F): int + function FP(g: G): int + function FR(h: H): int + + method M(x: X, y: Y, z: Z) + { + Q(x); + P(x); + R(x); + Q(y); // error: auto-init mismatch + P(y); + R(y); + Q(z); // error: auto-init mismatch + P(z); // error: auto-init mismatch + R(z); + } + + method N(x: X, y: Y, z: Z) returns (u: int) + { + u := FQ(x); + u := FP(x); + u := FR(x); + u := FQ(y); // error: auto-init mismatch + u := FP(y); + u := FR(y); + u := FQ(z); // error: auto-init mismatch + u := FP(z); // error: auto-init mismatch + u := FR(z); + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy.expect new file mode 100644 index 00000000000..a937bf71b38 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors6.dfy.expect @@ -0,0 +1,51 @@ +ResolutionErrors6.dfy(31,26): Error: type parameter (G) passed to method GimmieOne must support auto-initialization (got Yt) +ResolutionErrors6.dfy(44,50): Error: type of the receiver is not fully determined at this program point +ResolutionErrors6.dfy(71,11): Error: name of type (Cache) is used as a variable +ResolutionErrors6.dfy(71,17): Error: incorrect type for selection into ? (got X) +ResolutionErrors6.dfy(79,13): Warning: the quantifier has the form 'exists x :: A ==> B', which most often is a typo for 'exists x :: A && B'; if you think otherwise, rewrite as 'exists x :: (A ==> B)' or 'exists x :: !A || B' to suppress this warning +ResolutionErrors6.dfy(89,13): Warning: the quantifier has the form 'exists x :: A ==> B', which most often is a typo for 'exists x :: A && B'; if you think otherwise, rewrite as 'exists x :: (A ==> B)' or 'exists x :: !A || B' to suppress this warning +ResolutionErrors6.dfy(123,13): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(124,13): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(126,11): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(147,4): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(157,15): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(158,22): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(164,15): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(165,22): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors6.dfy(182,23): Error: type of left argument to * (int) must agree with the result type (bool) +ResolutionErrors6.dfy(182,23): Error: type of * must be of a numeric type, bitvector type, or a set-like type (instead got bool) +ResolutionErrors6.dfy(181,13): Error: not resolving module 'V' because there were errors in resolving its nested module 'W' +ResolutionErrors6.dfy(191,21): Error: type of left argument to * (int) must agree with the result type (bool) +ResolutionErrors6.dfy(191,21): Error: type of * must be of a numeric type, bitvector type, or a set-like type (instead got bool) +ResolutionErrors6.dfy(190,11): Error: not resolving module 'V' because there were errors in resolving its nested module 'W' +ResolutionErrors6.dfy(177,9): Error: not resolving module 'ErrorsFromNestedModules' because there were errors in resolving its nested module 'U' +ResolutionErrors6.dfy(199,11): Error: duplicate name of top-level declaration: G +ResolutionErrors6.dfy(202,10): Related location +ResolutionErrors6.dfy(205,11): Error: duplicate name of top-level declaration: H +ResolutionErrors6.dfy(203,10): Related location +ResolutionErrors6.dfy(221,4): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(223,20): Error: ghost variables such as y are allowed only in specification contexts. y was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(232,6): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(233,4): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(242,6): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(243,8): Error: ghost variables such as i are allowed only in specification contexts. i was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(244,4): Error: ghost variables such as b are allowed only in specification contexts. b was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(256,13): Error: expression is not allowed to refer to constant field x +ResolutionErrors6.dfy(258,22): Error: expression is not allowed to refer to constant field x +ResolutionErrors6.dfy(269,13): Error: arguments must be of a set or multiset type (got map) +ResolutionErrors6.dfy(278,4): Error: expect statement is not allowed in this context (because this is a ghost method or because the statement is guarded by a specification-only expression) +ResolutionErrors6.dfy(284,11): Error: quantifiers in non-ghost contexts must be compilable, but Dafny's heuristics can't figure out how to produce or compile a bounded set of values for 'i' +ResolutionErrors6.dfy(285,21): Error: ghost variables such as g are allowed only in specification contexts. g was inferred to be ghost based on its declaration or initialization. +ResolutionErrors6.dfy(297,15): Error: RHS (of type X) not assignable to LHS (of type X) +ResolutionErrors6.dfy(302,10): Error: type of corresponding source/RHS (X) does not match type of bound variable (X) +ResolutionErrors6.dfy(330,16): Error: sequence update requires the value to have the element type of the sequence (got Trait) +ResolutionErrors6.dfy(335,11): Error: multiset update requires domain element to be of type Elem (got Trait) +ResolutionErrors6.dfy(341,13): Error: map update requires domain element to be of type Elem (got Trait) +ResolutionErrors6.dfy(343,18): Error: map update requires the value to have the range type Elem (got Trait) +ResolutionErrors6.dfy(368,5): Error: type parameter (F) passed to method Q must support auto-initialization (got Y) (perhaps try declaring type parameter 'Y' on line 363 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors6.dfy(371,5): Error: type parameter (F) passed to method Q must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 363 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors6.dfy(372,5): Error: type parameter (G) passed to method P must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 363 as 'Z(00)', which says it can only be instantiated with a nonempty type) +ResolutionErrors6.dfy(381,9): Error: type parameter (F) passed to function FQ must support auto-initialization (got Y) (perhaps try declaring type parameter 'Y' on line 376 as 'Y(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors6.dfy(384,9): Error: type parameter (F) passed to function FQ must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 376 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors6.dfy(385,9): Error: type parameter (G) passed to function FP must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 376 as 'Z(00)', which says it can only be instantiated with a nonempty type) +46 resolution/type errors detected in ResolutionErrors6.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy new file mode 100644 index 00000000000..6ef93602379 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy @@ -0,0 +1,368 @@ +// RUN: %exits-with 2 %verify "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// --------------- ghost function error messages ------------------------------ + +module GhostFunctionErrorMessages { + ghost function GhostFunction(): int + ghost predicate GhostPredicate() + least predicate LeastPredicate() + greatest predicate GreatestPredicate() + twostate function TwoFunction(): int + twostate predicate TwoPredicate() + + method GhostsUsedInCompiledContexts() { + var x, b; + x := GhostFunction(); // error + b := GhostPredicate(); // error + b := LeastPredicate(); // error + b := GreatestPredicate(); // error + x := TwoFunction(); // error + b := TwoPredicate(); // error + } +} + +module TypeParameterCount { + ghost function F0(): int + ghost function F1(): int + ghost function F2(): int + method M0() + method M1() + method M2() + + method TestFunction() { + var x; + x := F0(); + x := F1(); // type argument inferred + x := F2(); // type arguments inferred + x := F0(); // error: wrong number of type parameters + x := F1(); + x := F2(); // error: wrong number of type parameters + x := F0(); // error: wrong number of type parameters + x := F1(); // error: wrong number of type parameters + x := F2(); + } + + method TestMethods() { + M0(); + M1(); // type argument inferred + M2(); // type arguments inferred + M0(); // error: wrong number of type parameters + M1(); + M2(); // error: wrong number of type parameters + M0(); // error: wrong number of type parameters + M1(); // error: wrong number of type parameters + M2(); + } +} + +module AutoGhostRegressions { + datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) + + method Test() { + var q := Quad(20, 30, 40, 50); + print q, "\n"; + + var Quad(a, b, c, d) := q; + print c, "\n"; // error: c is ghost (this was once not handled correctly) + + match q { + case Quad(r, s, t, u) => + print t, "\n"; // error: t is ghost + } + + ghost var p := Quad(20, 30, 40, 50); + var Quad(a', b', c', d') := p; + print a', "\n"; // error: a' is ghost + print c', "\n"; // error: c' is ghost + } + + datatype NoEquality = NoEquality(ghost u: int) + newtype NT = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context + type ST = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context +} + +module TypeCharacteristicsInGhostCode { + function MustBeNonempty(): int { 5 } + function MustBeAutoInit(): int { 5 } + function MustSupportEquality(): int { 5 } + function NoReferences(): int { 5 } + + type PossiblyEmpty = x: int | true witness * + type Nonempty = x: int | true ghost witness 0 + datatype NoEquality = NoEquality(ghost u: int) + class Class { } + type Good = bool + + method TestCompiled() + { + var w; + + w := MustBeNonempty(); // error + w := MustBeNonempty(); + w := MustBeNonempty(); + w := MustBeNonempty(); + w := MustBeNonempty(); + w := MustBeNonempty(); // error (a hint is given) + + w := MustBeAutoInit(); // error + w := MustBeAutoInit(); // error + w := MustBeAutoInit(); + w := MustBeAutoInit(); + w := MustBeAutoInit(); + w := MustBeAutoInit(); // error (a hint is given) + + w := MustSupportEquality(); + w := MustSupportEquality(); + w := MustSupportEquality(); // error + w := MustSupportEquality(); + w := MustSupportEquality(); + w := MustSupportEquality(); // error (a hint is given) + + w := NoReferences(); + w := NoReferences(); + w := NoReferences(); + w := NoReferences(); // error + w := NoReferences(); + w := NoReferences(); // error (a hint is given) + } + + method TestGhost() + { + ghost var w; + + w := MustBeNonempty(); // error + w := MustBeNonempty(); + w := MustBeNonempty(); + w := MustBeNonempty(); + w := MustBeNonempty(); + + w := MustBeAutoInit(); // error + w := MustBeAutoInit(); // fine, because the call appears in a ghost context + w := MustBeAutoInit(); + w := MustBeAutoInit(); + w := MustBeAutoInit(); + + w := MustSupportEquality(); + w := MustSupportEquality(); + w := MustSupportEquality(); + w := MustSupportEquality(); + w := MustSupportEquality(); + + w := NoReferences(); + w := NoReferences(); + w := NoReferences(); + w := NoReferences(); // error + w := NoReferences(); + } + + function FF(a: bool, ghost b: bool): int { 5 } + method MM(a: bool, ghost b: bool) { } + function GetInt(): int { 2 } + + method GhostContexts(x: T, y: T) { + var r; + r := FF(x == y, true); // error: T must support equality + r := FF(true, x == y); // no problem, since this is a ghost context + MM(x == y, true); // error: T must support equality + MM(true, x == y); // no problem, since this is a ghost context + + r := FF(GetInt() == 0, true); // error: type argument must support equality + r := FF(true, GetInt() == 0); // okay, since this is a ghost context + MM(GetInt() == 0, true); // error: type argument must support equality + MM(true, GetInt() == 0); // okay, since this is a ghost context + + var q0 := Quad(GetInt() == 0, true, true, true); // error: type argument must support equality + var q1 := Quad(true, true, GetInt() == 0, true); // fine, since in a ghost context + } + + datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) + datatype QuadEq = QuadEq(0: T, 1: T, ghost 2: U, ghost 3: U) + + method VarDecls(x: T, y: T) { + var a := x == y; // error: this requires T to support equality + ghost var b := x == y; // fine + + var q := Quad(x, y, x, y); + var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost + var c := k == l; // error: this requires T to support equality + var d := m == n; // fine, since d is implicitly ghost + d, c, d := m == n, k == l, m == n; // error: "k == l" requires T to support equality + + var q' := QuadEq([x], [x], [0], [0]); // error: seq requires T to support equality + var q'' := QuadEq([0], [0], [x], [x]); // error: seq requires T to support equality + } + + newtype NT = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context + type ST = x | var s: set := {}; |s| <= x // fine, since constraint is a ghost context + + method LetVarDecls(x: T, y: T) { + var lhs; + lhs := + var a := x == y; // error: this requires T to support equality + 0; + lhs := + ghost var b := x == y; // fine + 0; + + var q := Quad(x, y, x, y); + var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost + lhs := + var c := k == l; // error: this requires T to support equality + 0; + lhs := + var d := m == n; // fine, since d is implicitly ghost + 0; + + ghost var ghostLhs; + ghostLhs := + var a := x == y; // fine + 0; + ghostLhs := + ghost var b := x == y; // fine + 0; + } + + datatype DatatypeHasMembersToo = Abc | Def { + method Test() { + var w; + w := MustBeNonempty(); // error + w := MustBeAutoInit(); // error + w := MustBeAutoInit(); // error + w := MustSupportEquality(); // error + w := NoReferences(); // error + } + } + + newtype NewtypeHasMembersToo = x: int | x == MustBeNonempty() // error: constraint has bad type instantiation + witness MustBeNonempty() // error: witness expression has bad type instantiation + { + method Test() { + var w; + w := MustBeNonempty(); // error + w := MustBeAutoInit(); // error + w := MustBeAutoInit(); // error + w := MustSupportEquality(); // error + w := NoReferences(); // error + } + } + + type SubsetTypeHasExpressionToo = x: int | x == MustBeNonempty() // error: constraint has bad type instantiation + witness MustBeNonempty() // error: witness expression has bad type instantiation + + newtype NT_CompiledWitness = x | 0 <= x + witness MustSupportEquality() // error + newtype NT_GhostWitness = x | 0 <= x + ghost witness MustSupportEquality() // fine, since it's ghost + type ST_CompiledWitness = x | 0 <= x + witness MustSupportEquality() // error + type ST_GhostWitness = x | 0 <= x + ghost witness MustSupportEquality() // fine, since it's ghost + + trait + {:MyAttribute MustSupportEquality(), MustBeNonempty()} // error: about MustBeNonempty (no prob with (==)) + MyTrait + { + const x := (CallMyLemma(MustSupportEquality(), MustBeNonempty()); 23) // error: about MustBeNonempty (no prob with (==)) + } + lemma CallMyLemma(x: int, y: int) +} + +module MoreAutoGhostTests { + datatype Quad = Quad(0: T, 1: T, ghost 2: U, ghost 3: U) + + method SomeLetVarsAreGhost0(q: Quad) returns (r: int) { + r := + var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost + k + l; + } + + method SomeLetVarsAreGhost1(q: Quad) returns (r: int) { + r := + var Quad(k, l, m, n) := q; // k,l are compiled; m,n are ghost + m; // error: m is ghost + } + + method AssignSuchThat(ghost m: int) { + var d :| d == m; // error: LHS is not inferred to be ghost for :| + } + + function LetSuchThat(ghost m: int): int { + var d :| d == m; // error: LHS is not inferred to be ghost for :| + 0 + } +} + +module RelaxedAutoInitChecking { + // In a ghost context, the (==) is never relevant. Therefore, checking of adherence to (==) for + // type arguments is suppressed in ghost contexts. + // Similarly, in a ghost context, there's no difference between (0) and (00). Therefore, a + // formal parameter that expects (0) can take either a (0) or a (00) in a ghost context. + + function MustBeNonempty(): int { 5 } + function MustBeAutoInit(): int { 5 } + function MustSupportEquality(): int { 5 } + function NoReferences(): int { 5 } + + type PossiblyEmpty = x: int | true witness * + type Nonempty = x: int | true ghost witness 0 + datatype NoEquality = NoEquality(ghost u: int) + class Class { } + type Good = bool + + method M(compiledValue: int, ghost ghostValue: int) + + method TestCompiled() + { + M(MustBeNonempty(), 0); // error + M(MustBeNonempty(), 0); + M(MustBeNonempty(), 0); + M(MustBeNonempty(), 0); + M(MustBeNonempty(), 0); + + M(MustBeAutoInit(), 0); // error + M(MustBeAutoInit(), 0); // error + M(MustBeAutoInit(), 0); + M(MustBeAutoInit(), 0); + M(MustBeAutoInit(), 0); + + M(MustSupportEquality(), 0); + M(MustSupportEquality(), 0); + M(MustSupportEquality(), 0); // error + M(MustSupportEquality(), 0); + M(MustSupportEquality(), 0); + + M(NoReferences(), 0); + M(NoReferences(), 0); + M(NoReferences(), 0); + M(NoReferences(), 0); // error + M(NoReferences(), 0); + } + + method TestGhost() + { + M(0, MustBeNonempty()); // error + M(0, MustBeNonempty()); + M(0, MustBeNonempty()); + M(0, MustBeNonempty()); + M(0, MustBeNonempty()); + + M(0, MustBeAutoInit()); // error + M(0, MustBeAutoInit()); // this is fine in a ghost context + M(0, MustBeAutoInit()); + M(0, MustBeAutoInit()); + M(0, MustBeAutoInit()); + + M(0, MustSupportEquality()); + M(0, MustSupportEquality()); + M(0, MustSupportEquality()); // this is fine in a ghost context + M(0, MustSupportEquality()); + M(0, MustSupportEquality()); + + M(0, NoReferences()); + M(0, NoReferences()); + M(0, NoReferences()); + M(0, NoReferences()); // error + M(0, NoReferences()); + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy.expect new file mode 100644 index 00000000000..45ec14b79f3 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors7.dfy.expect @@ -0,0 +1,72 @@ +ResolutionErrors7.dfy(16,9): Error: a call to a ghost function is allowed only in specification contexts (consider declaring the function without the 'ghost' keyword) +ResolutionErrors7.dfy(17,9): Error: a call to a ghost predicate is allowed only in specification contexts (consider declaring the predicate without the 'ghost' keyword) +ResolutionErrors7.dfy(18,9): Error: a call to a least predicate is allowed only in specification contexts +ResolutionErrors7.dfy(19,9): Error: a call to a greatest predicate is allowed only in specification contexts +ResolutionErrors7.dfy(20,9): Error: a call to a twostate function is allowed only in specification contexts +ResolutionErrors7.dfy(21,9): Error: a call to a twostate predicate is allowed only in specification contexts +ResolutionErrors7.dfy(38,9): Error: function 'F0' expects 0 type arguments (got 1) +ResolutionErrors7.dfy(40,9): Error: function 'F2' expects 2 type arguments (got 1) +ResolutionErrors7.dfy(41,9): Error: function 'F0' expects 0 type arguments (got 2) +ResolutionErrors7.dfy(42,9): Error: function 'F1' expects 1 type argument (got 2) +ResolutionErrors7.dfy(50,4): Error: method 'M0' expects 0 type arguments (got 1) +ResolutionErrors7.dfy(52,4): Error: method 'M2' expects 2 type arguments (got 1) +ResolutionErrors7.dfy(53,4): Error: method 'M0' expects 0 type arguments (got 2) +ResolutionErrors7.dfy(54,4): Error: method 'M1' expects 1 type argument (got 2) +ResolutionErrors7.dfy(67,10): Error: ghost variables such as c are allowed only in specification contexts. c was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(71,14): Error: ghost variables such as t are allowed only in specification contexts. t was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(76,10): Error: ghost variables such as a' are allowed only in specification contexts. a' was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(77,10): Error: ghost variables such as c' are allowed only in specification contexts. c' was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(101,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(106,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got Z) (perhaps try declaring type parameter 'Z' on line 97 as 'Z(00)', which says it can only be instantiated with a nonempty type) +ResolutionErrors7.dfy(108,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(109,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) +ResolutionErrors7.dfy(113,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Z) (perhaps try declaring type parameter 'Z' on line 97 as 'Z(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors7.dfy(117,9): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(120,9): Error: type parameter (T) passed to function MustSupportEquality must support equality (got Z) (perhaps try declaring type parameter 'Z' on line 97 as 'Z(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(125,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +ResolutionErrors7.dfy(127,9): Error: type parameter (T) passed to function NoReferences must support no references (got Z) (perhaps try declaring type parameter 'Z' on line 97 as 'Z(!new)', which says it can only be instantiated with a type that contains no references) +ResolutionErrors7.dfy(134,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(140,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(155,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +ResolutionErrors7.dfy(165,12): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 163 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(167,7): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 163 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(170,12): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) +ResolutionErrors7.dfy(172,7): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) +ResolutionErrors7.dfy(175,19): Error: type parameter (T) passed to function GetInt must support equality (got NoEquality) +ResolutionErrors7.dfy(183,13): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 182 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(188,13): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 182 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(190,23): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 182 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(192,14): Error: type parameter 0 (T) passed to type QuadEq must support equality (got seq) +ResolutionErrors7.dfy(193,15): Error: type parameter 1 (U) passed to type QuadEq must support equality (got seq) +ResolutionErrors7.dfy(202,15): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 199 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(211,15): Error: == can only be applied to expressions of types that support equality (got T) (perhaps try declaring type parameter 'T' on line 199 as 'T(==)', which says it can only be instantiated with a type that supports equality) +ResolutionErrors7.dfy(229,11): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(230,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(231,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) +ResolutionErrors7.dfy(232,11): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(233,11): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +ResolutionErrors7.dfy(237,47): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(238,12): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(242,11): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(243,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(244,11): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) +ResolutionErrors7.dfy(245,11): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(246,11): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +ResolutionErrors7.dfy(250,50): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(251,12): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(254,12): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(258,12): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(263,53): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(266,63): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(283,6): Error: ghost variables such as m are allowed only in specification contexts. m was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(287,8): Error: non-ghost variable cannot be assigned a value that depends on a ghost +ResolutionErrors7.dfy(291,18): Error: ghost variables such as m are allowed only in specification contexts. m was inferred to be ghost based on its declaration or initialization. +ResolutionErrors7.dfy(317,6): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(323,6): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(324,6): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got Nonempty) +ResolutionErrors7.dfy(331,6): Error: type parameter (T) passed to function MustSupportEquality must support equality (got NoEquality) +ResolutionErrors7.dfy(338,6): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +ResolutionErrors7.dfy(344,9): Error: type parameter (T) passed to function MustBeNonempty must support nonempty (got PossiblyEmpty) +ResolutionErrors7.dfy(350,9): Error: type parameter (T) passed to function MustBeAutoInit must support auto-initialization (got PossiblyEmpty) +ResolutionErrors7.dfy(365,9): Error: type parameter (T) passed to function NoReferences must support no references (got Class?) +71 resolution/type errors detected in ResolutionErrors7.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy new file mode 100644 index 00000000000..452110e9b20 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy @@ -0,0 +1,462 @@ +// RUN: %exits-with 2 %verify "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// --------------- let-such-that ghost regressions ------------------------------ + +module LetSuchThatGhost { + ghost predicate True(t: T) { true } + + function F(s: set): int + requires s != {} + { + // once, the RHS for p was (bogusly) considered ghost, which made p ghost, + // which made this body illegal + var p := + var e :| e in s; + true; + if p then 6 else 8 + } + + function G(s: set): int + requires s != {} + { + // again, e and p are both non-ghost + var p := + var e :| e in s; + e == e; + if p then 6 else 8 + } + + function H(s: set): int + requires s != {} + { + // here, e is ghost, but p is still not + var p := + var e :| e in s && True(e); + true; + if p then 6 else 8 + } + + function I(s: set): int + requires s != {} + { + // here, e is ghost, and therefore so is p + var p := + var e :| e in s && True(e); + e == e; + if p then 6 else 8 // error: p is ghost + } +} + +// --------------- hint restrictions in non-while loops ------------------------------ + +module HintRestrictionsOtherLoops { + class C { + ghost function F(): int + { + calc { + 6; + { var x := 8; + while + modifies this // error: cannot use a modifies clause on a loop inside a hint + { + case x != 0 => x := x - 1; + } + } + 6; + { for i := 0 to 8 + modifies this // error: cannot use a modifies clause on a loop inside a hint + { + } + } + 6; + } + 5 + } + } +} + +// --------------- regressions: types of arguments to fresh, unchanged, modifies, reads ------------------------------ + +module FrameTypes { + // ----- fresh takes an expression as its argument + + method FreshArgumentType0( + o: object, + s: set, ss: iset, + q: seq, + ms: multiset, + mp: map, mp2: map, + im: imap) + { + ghost var b; + b := fresh(o); + b := fresh(s); + b := fresh(ss); + b := fresh(q); + b := fresh(ms); // error: wrong argument type for fresh + b := fresh(mp); // error: wrong argument type for fresh + b := fresh(mp2); // error: wrong argument type for fresh + b := fresh(im); // error: wrong argument type for fresh + } + + method FreshArgumentType1(x: int, s: set, ss: iset, q: seq) { + ghost var b; + b := fresh(x); // error: wrong argument type for fresh + b := fresh(s); // error: wrong argument type for fresh + b := fresh(ss); // error: wrong argument type for fresh + b := fresh(q); // error: wrong argument type for fresh + } + + method FreshArgumentType2(f: int -> int, g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int) { + ghost var b; + b := fresh(f); // error: wrong argument type for fresh + b := fresh(g); // error: wrong argument type for fresh + b := fresh(h); // error: wrong argument type for fresh + b := fresh(j); // error: wrong argument type for fresh + b := fresh(k); // error: wrong argument type for fresh + } + + // ----- unchanged, modifies, and reads take frame expressions as their arguments + + method UnchangedArgumentType0( + o: object, + s: set, ss: iset, + q: seq, + ms: multiset, + mp: map, mp2: map, + im: imap) + { + ghost var b; + b := unchanged(o); + b := unchanged(s); + b := unchanged(ss); + b := unchanged(q); + b := unchanged(ms); + b := unchanged(mp); // error: wrong argument type for unchanged + b := unchanged(mp2); // error: wrong argument type for unchanged + b := unchanged(im); // error: wrong argument type for unchanged + } + + method UnchangedArgumentType1(x: int, s: set, ss: iset, q: seq) { + ghost var b; + b := unchanged(x); // error: wrong argument type for unchanged + b := unchanged(s); // error: wrong argument type for unchanged + b := unchanged(ss); // error: wrong argument type for unchanged + b := unchanged(q); // error: wrong argument type for unchanged + } + + method UnchangedArgumentType2( + f: int -> int, + g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, + l: bool -> multiset, m: bool -> map) + { + ghost var b; + b := unchanged(f); // error: wrong argument type for unchanged + b := unchanged(g); // error: wrong argument type for unchanged + b := unchanged(h); // error: wrong argument type for unchanged + b := unchanged(i); // error: wrong argument type for unchanged + b := unchanged(j); // error: wrong argument type for unchanged + b := unchanged(k); // error: wrong argument type for unchanged + b := unchanged(l); // error: wrong argument type for unchanged + b := unchanged(m); // error: wrong argument type for unchanged + } + + method ModifiesArgumentType0( + o: object, + s: set, ss: iset, + q: seq, + ms: multiset, + mp: map, mp2: map, + im: imap) + modifies o + modifies s + modifies ss + modifies q + modifies ms + modifies mp // error: wrong argument type for modifies + modifies mp2 // error: wrong argument type for modifies + modifies im // error: wrong argument type for modifies + { + } + + method ModifiesArgumentType1(x: int, s: set, ss: iset, q: seq) + modifies x // error: wrong argument type for modifies + modifies s // error: wrong argument type for modifies + modifies ss // error: wrong argument type for modifies + modifies q // error: wrong argument type for modifies + { + } + + method ModifiesArgumentType2( + f: int -> int, + g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, + l: bool -> multiset, m: bool -> map) + modifies f // error: wrong argument type for modifies + modifies g // error: wrong argument type for modifies + modifies h // error: wrong argument type for modifies + modifies i // error: wrong argument type for modifies + modifies j // error: wrong argument type for modifies + modifies k // error: wrong argument type for modifies + modifies l // error: wrong argument type for modifies + modifies m // error: wrong argument type for modifies + { + } + + ghost predicate ReadsArgumentType0( + o: object, + s: set, ss: iset, + q: seq, + ms: multiset, + mp: map, mp2: map, + im: imap) + reads o + reads s + reads ss + reads q + reads ms + reads mp // error: wrong argument type for reads + reads mp2 // error: wrong argument type for reads + reads im // error: wrong argument type for reads + { + true + } + + ghost predicate ReadsArgumentType1(x: int, s: set, ss: iset, q: seq) + reads x // error: wrong argument type for reads + reads s // error: wrong argument type for reads + reads ss // error: wrong argument type for reads + reads q // error: wrong argument type for reads + { + true + } + + ghost predicate ReadsArgumentType2( + f: int -> int, + g: int -> object, h: int -> set, i: int -> iset, j: int -> seq, k: set -> int, + l: bool -> multiset, m: bool -> map) + reads f // error: wrong argument type for reads + reads g // error: a function must be to a collection of references + reads h + reads i + reads j + reads k // error: wrong argument type for reads + reads l + reads m // error: wrong argument type for reads + { + true + } +} + +module Continue0 { + method BadTargetsLevels(a: int, b: int, c: int) { + for i := 0 to 100 { + for j := 0 to 100 { + for k := 0 to 100 { + if + case k == a => + continue; + case k == b => + break continue; + case k == c => + break break continue; + case k == a + b + c => + break break break continue; // error: too many levels + } + } + } + } + + method BadTargetsLabels(a: int, b: int, c: int) { + label A: + for i := 0 to 100 { + label B0: label B1: + for j := 0 to 100 { + label C: + for k := 0 to 100 { + if + case k == a => + continue C; + case k == b => + continue B0; + case k == b => + continue B1; + case k == c => + continue A; + } + } + } + } + + method NonLoopLabels(a: int, b: int, c: int) { + // the following labels are attached to BlockStmt's, not loops + label X: { + for i := 0 to 100 { + label Y0: label Y1: { + for j := 0 to 100 { + label Z: { + for k := 0 to 100 { + if + case k == a => + continue X; // error: X is not a loop label + case k == b => + continue Y0; // error: Y0 is not a loop label + case k == b => + continue Y1; // error: Y1 is not a loop label + case k == c => + continue Z; // error: Z is not a loop label + } + } + } + } + } + } + } + + method SimpleBadJumps0() { + break; // error: cannot "break" from here + } + + method SimpleBadJumps1() { + continue; // error: cannot "continue" from here + } + + method SimpleBadJumps2() { + label X: { + if + case true => break; // error: cannot "break" from here + case true => continue; // error: cannot "continue" from here + case true => break X; + case true => continue X; // error: X is not a loop label + } + } + + method GhostContinueAssertBy(ghost t: int, ghost u: nat) + { + label L: + for i := 0 to 100 { + assert true by { + for j := 0 to 100 { + if j == t { + break; + } else if j == u { + continue; + } + } + if + case true => break; // error: cannot jump outside the assert-by + case true => continue; // error: cannot jump outside the assert-by + case true => break L; // error: cannot jump outside the assert-by + case true => continue L; // error: cannot jump outside the assert-by + } + } + } +} + +module Continue1 { + method GhostContinueLevels(ghost t: int, ghost u: nat) + { + var m := 0; + for i := 0 to 100 { + if i == t { + // The following "continue" would pass the increment to m + continue; // error: continue from ghost context must target a ghost loop + } + m := m + 1; + } + + for i := 0 to 100 { + m := m + 1; + // The following "break" would potentially pass both increments to m + if i == t { + break; // error: break from ghost context must target a ghost loop + } + m := m + 1; + } + + for i := 0 to 100 { + if i == t { + // Even though there's no statement in the loop body after this ghost if, the continue violates the rule + continue; // error: continue from ghost context must target a ghost loop + } + } + + for i := 0 to 100 { + for j := 0 to u { + if i == t { + continue; // fine + } + } + } + + for i := 0 to 100 { + for j := 0 to u { + if i == t { + break continue; // error: continue from ghost context must target a ghost loop + } + } + } + + for i := 0 to 100 + u { + for j := 0 to u { + if i == t { + break continue; // fine + } + } + } + } + + method GhostContinueLabels(ghost t: int, ghost u: nat) + { + label Outer: + for i := 0 to 100 { + label Inner: + for j := 0 to u { + if j == t { + continue Inner; // fine + } else if j == 20 + t { + continue Outer; // error: continue from ghost context must target a ghost loop + } + } + } + } +} + +module LabelRegressions { + // The cases of if-case, while-case, and match statements are List's, which are essentially + // a BlockStmt but without the curly braces. Each Statement in such a List can have labels, so + // it's important to ResolveStatementWithLabels, not ResolveStatement. Alas, that was once not the + // case (pun intended). + // There's also something analogous going on in the Verifier, where lists of statements should call + // TrStmtList, not just call TrStmt on every Statement in the List. (See method LabelRegressions() + // in Test/comp/ForLoops-Compilation.dfy.) + method IfCaseRegression() { + if + case true => + label Loop: + for k := 0 to 10 { + continue Loop; + break Loop; + } + } + + method WhileCaseRegression() { + while + case true => + label Loop: + for k := 0 to 10 { + continue Loop; + break Loop; + } + } + + method Match() { + match (0, 0) + case (_, _) => + label Loop: + for k := 0 to 10 { + break Loop; + continue Loop; + } + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy.expect new file mode 100644 index 00000000000..9f4af2e8871 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors8.dfy.expect @@ -0,0 +1,77 @@ +ResolutionErrors8.dfy(47,7): Error: ghost variables such as p are allowed only in specification contexts. p was inferred to be ghost based on its declaration or initialization. +ResolutionErrors8.dfy(61,21): Error: a loop in a hint is not allowed to use 'modifies' clauses +ResolutionErrors8.dfy(68,21): Error: a loop in a hint is not allowed to use 'modifies' clauses +ResolutionErrors8.dfy(97,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got multiset) +ResolutionErrors8.dfy(98,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got map) +ResolutionErrors8.dfy(99,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got map) +ResolutionErrors8.dfy(100,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got imap) +ResolutionErrors8.dfy(105,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int) +ResolutionErrors8.dfy(106,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got set) +ResolutionErrors8.dfy(107,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got iset) +ResolutionErrors8.dfy(108,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got seq) +ResolutionErrors8.dfy(113,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> int) +ResolutionErrors8.dfy(114,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> object) +ResolutionErrors8.dfy(115,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> set) +ResolutionErrors8.dfy(116,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got int -> seq) +ResolutionErrors8.dfy(117,9): Error: the argument of a fresh expression must denote an object or a set or sequence of objects (instead got set -> int) +ResolutionErrors8.dfy(136,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(137,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(138,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got imap) +ResolutionErrors8.dfy(143,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int) +ResolutionErrors8.dfy(144,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got set) +ResolutionErrors8.dfy(145,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got iset) +ResolutionErrors8.dfy(146,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got seq) +ResolutionErrors8.dfy(155,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> int) +ResolutionErrors8.dfy(156,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> object) +ResolutionErrors8.dfy(157,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> set) +ResolutionErrors8.dfy(158,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> iset) +ResolutionErrors8.dfy(159,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> seq) +ResolutionErrors8.dfy(160,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got set -> int) +ResolutionErrors8.dfy(161,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> multiset) +ResolutionErrors8.dfy(162,19): Error: an unchanged expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> map) +ResolutionErrors8.dfy(177,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(178,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(179,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got imap) +ResolutionErrors8.dfy(184,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int) +ResolutionErrors8.dfy(185,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got set) +ResolutionErrors8.dfy(186,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got iset) +ResolutionErrors8.dfy(187,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got seq) +ResolutionErrors8.dfy(195,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> int) +ResolutionErrors8.dfy(196,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> object) +ResolutionErrors8.dfy(197,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> set) +ResolutionErrors8.dfy(198,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> iset) +ResolutionErrors8.dfy(199,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got int -> seq) +ResolutionErrors8.dfy(200,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got set -> int) +ResolutionErrors8.dfy(201,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> multiset) +ResolutionErrors8.dfy(202,13): Error: a modifies-clause expression must denote an object or a set/iset/multiset/seq of objects (instead got bool -> map) +ResolutionErrors8.dfy(218,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(219,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got map) +ResolutionErrors8.dfy(220,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got imap) +ResolutionErrors8.dfy(226,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int) +ResolutionErrors8.dfy(227,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got set) +ResolutionErrors8.dfy(228,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got iset) +ResolutionErrors8.dfy(229,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got seq) +ResolutionErrors8.dfy(238,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int -> int) +ResolutionErrors8.dfy(239,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got int -> object) +ResolutionErrors8.dfy(243,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got set -> int) +ResolutionErrors8.dfy(245,10): Error: a reads-clause expression must denote an object, a set/iset/multiset/seq of objects, or a function to a set/iset/multiset/seq of objects (instead got bool -> map) +ResolutionErrors8.dfy(264,12): Error: a 'break break break continue' statement is allowed only in contexts with 4 enclosing loops, but the current context only has 3 +ResolutionErrors8.dfy(301,27): Error: continue label must designate a loop: X +ResolutionErrors8.dfy(303,27): Error: continue label must designate a loop: Y0 +ResolutionErrors8.dfy(305,27): Error: continue label must designate a loop: Y1 +ResolutionErrors8.dfy(307,27): Error: continue label must designate a loop: Z +ResolutionErrors8.dfy(317,4): Error: a non-labeled 'break' statement is allowed only in loops +ResolutionErrors8.dfy(321,4): Error: a non-labeled 'continue' statement is allowed only in loops +ResolutionErrors8.dfy(327,19): Error: a non-labeled 'break' statement is allowed only in loops +ResolutionErrors8.dfy(328,19): Error: a non-labeled 'continue' statement is allowed only in loops +ResolutionErrors8.dfy(330,28): Error: continue label must designate a loop: X +ResolutionErrors8.dfy(347,21): Error: a non-labeled 'break' statement is allowed only in loops +ResolutionErrors8.dfy(348,21): Error: a non-labeled 'continue' statement is allowed only in loops +ResolutionErrors8.dfy(349,27): Error: break label is undefined or not in scope: L +ResolutionErrors8.dfy(350,30): Error: continue label is undefined or not in scope: L +ResolutionErrors8.dfy(363,8): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop +ResolutionErrors8.dfy(372,8): Error: ghost-context break statement is not allowed to break out of non-ghost loop +ResolutionErrors8.dfy(380,8): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop +ResolutionErrors8.dfy(395,10): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop +ResolutionErrors8.dfy(418,10): Error: ghost-context continue statement is not allowed to continue out of non-ghost loop +76 resolution/type errors detected in ResolutionErrors8.dfy diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy new file mode 100644 index 00000000000..966e099ed22 --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy @@ -0,0 +1,219 @@ +// RUN: %exits-with 2 %verify --allow-axioms "%s" > "%t" +// RUN: %diff "%s.expect" "%t" + +// --------------- regressions: using "this" in places where there is no enclosing class/type ------------------------------ + +module UseOfThis { + // The following uses of "this." once caused the resolver to crash. + + type {:badUseOfThis this.K} OpaqueType { // error: cannot use "this" here + const K: int + } + + newtype {:badUseOfThis this.x} Newtype = // error: cannot use "this" here + x: int | this.u // error: cannot use "this" here + witness this.u // error: cannot use "this" here + { + const K: int + } + + type {:badUseOfThis this.x} SynonymType = int // error: cannot use "this" here + + type {:badUseOfThis this.x} SubsetType = // error: cannot use "this" here + x: int | this.u // error: cannot use "this" here + witness this.u // error: cannot use "this" here + + trait {:badUseOfThis this.K} MyTrait { // error: cannot use "this" here + const K: int + } + + class {:badUseOfThis this.M} MyClass { // error: cannot use "this" here + const M: int + + var {:goodUseOfThis this.M} I: int + const {:goodUseOfThis this.M} J := 3 + method {:goodUseOfThis this.M} CM() + ensures {:goodUseOfThis this.M} true + ghost function {:goodUseOfThis this.M} CF(): int + ensures {:goodUseOfThis this.M} true + + static const {:badUseOfThis this.M} L := 3 // error: cannot use "this" here + static const N := this.M // error: cannot use "this" here + static method {:badUseOfThis this.M} SM() // error: cannot use "this" here + ensures {:badUseOfThis this.M} true // error: cannot use "this" here + static ghost function {:badUseOfThis this.M} SF(): int // error: cannot use "this" here + ensures {:badUseOfThis this.M} true // error: cannot use "this" here + } + + datatype Datatype = + | {:badUseOfThis this.K} DatatypeCtor // error: cannot use "this" here + { + const K: int + } +} + +module AutoInitTypeCheckRegression { + codatatype AutoStream = AutoNext(head: G, tail: AutoStream) + + function In(a: AutoStream): int // error: the argument to AutoStream is supposed to be an auto-init type + function Out(g: G): AutoStream // error: the argument to AutoStream is supposed to be an auto-init type + + method M(a: AutoStream) // error: the argument to AutoStream is supposed to be an auto-init type + method N(g: G) returns (a: AutoStream) // error: the argument to AutoStream is supposed to be an auto-init type +} + +module ClassConstructorTests { + class ClassWithConstructor { + var y: int + method NotTheOne() { } + constructor InitA() { } + constructor InitB() { y := 20; } + } + + class ClassWithoutConstructor { + method Init() modifies this { } + } + + method ConstructorTests() + { + var o := new object; // fine: does not have any constructors + + o := new ClassWithoutConstructor; // fine: don't need to call any particular method + o := new ClassWithoutConstructor.Init(); // this is also fine + + var c := new ClassWithConstructor.InitA(); + c := new ClassWithConstructor; // error: must call a constructor + c := new ClassWithConstructor.NotTheOne(); // error: must call a constructor, not an arbitrary method + c := new ClassWithConstructor.InitB(); + } + + class Y { + var data: int + constructor (x: int) + { + data := x; + } + method Test() { + var i := new Y(5); + i := new Y(7); + i := new Y; // error: the class has a constructor, so one must be used + var s := new Luci.Init(5); + s := new Luci.FromArray(null); + s := new Luci(false); + s := new Luci(true); + s := new Luci.M(); // error: there is a constructor, so one must be called + s := new Luci; // error: there is a constructor, so one must be called + var l := new Lamb; // fine, since there are no constructors (only methods) + l := new Lamb.Gwen(); + } + } + + class Luci { + constructor Init(y: int) { } + constructor (nameless: bool) { } + constructor FromArray(a: array) { } + method M() { } + } + + class Lamb { + method Jess() { } + method Gwen() { } + } +} + +module DefaultValues { + method Test() { + var r := 3 + MyFunction(7); + } + + function MyFunction(n: int, acc: int := n * false): int { // note, badly type default-value expression + 20 + } +} + +module IndexAdviceTypeRegressions { + type BV10 = x: bv10 | x != 999 witness 8 + + method Index(s: seq) + requires 2 < |s| + { + var k: bv10 := 2; + var a := s[k := "tjena"]; + + var K: BV10 := 2; + var A := s[K := "tjena"]; + + // rely on advice for the type + var i; + var b := s[i := "tjena"]; + } + + method Multi(matrix: array2) + requires matrix.Length0 == matrix.Length1 == 100 + { + var u := matrix[2, 5]; + var i := 2; + u := matrix[i, 5]; + var j := 2; + u := matrix[j, 5]; + j := j & j; + j := 16 as bv29; + // rely on advice for the type + var k; + u := matrix[k, 5]; + } + + method Array(arr: array) returns (r: real) + requires 10 <= arr.Length + { + var k: bv10 := 2; + r := arr[k]; + var K: BV10 := 2; + r := arr[K]; + } + + method Seq(s: seq) returns (r: real) + requires 10 <= |s| + { + var k: bv10 := 2; + r := s[k]; + var K: BV10 := 2; + r := s[K]; + } + + method Range(arr: array, s: seq) returns (r: seq) + requires 10 <= arr.Length <= |s| + { + var k: bv10 := 2; + var K: BV10 := 4; + r := arr[k..K]; + r := s[k..K]; + r := arr[k..]; + r := s[k..]; + r := arr[..K]; + r := s[..K]; + + // rely on advice for the type + var a, b, c, d; + r := arr[a..b]; + r := arr[c..]; + r := arr[..d]; + var w, x, y, z; + r := s[w..x]; + r := s[y..]; + r := s[..z]; + } + + method Bad(arr: array, s: seq, matrix: array2) returns (r: real, range: seq) { + r := arr[2.3]; // error: bad index type + r := s[2.3]; // error: bad index type + r := matrix[false, 2.3]; // error: bad index type + + range := arr[2.3..3.14]; // error (x2): bad index types + range := arr[2.3..]; // error: bad index type + range := arr[..3.14]; // error: bad index type + range := s[2.3..3.14]; // error (x2): bad index types + range := s[2.3..]; // error: bad index type + range := s[..3.14]; // error: bad index type + } +} diff --git a/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy.expect b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy.expect new file mode 100644 index 00000000000..a243ad4917b --- /dev/null +++ b/Source/IntegrationTests/TestFiles/LitTests/LitTest/dafny0/ResolutionErrors9.dfy.expect @@ -0,0 +1,59 @@ +ResolutionErrors9.dfy(14,13): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(14,18): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(23,13): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(23,18): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(41,22): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(9,22): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(9,27): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(15,12): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(15,17): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(13,25): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(13,30): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(20,22): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(20,27): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(24,12): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(24,17): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(22,22): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(22,27): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(26,23): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(26,28): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(40,32): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(43,29): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(42,33): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(45,29): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(44,41): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(30,23): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(30,28): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(49,21): Error: 'this' is not allowed in a 'static' context +ResolutionErrors9.dfy(49,26): Error: type of the receiver is not fully determined at this program point +ResolutionErrors9.dfy(58,20): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 58 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors9.dfy(59,25): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 59 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors9.dfy(61,17): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 61 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors9.dfy(62,32): Error: type parameter (G) passed to type AutoStream must support auto-initialization (got G) (perhaps try declaring type parameter 'G' on line 62 as 'G(0)', which says it can only be instantiated with a type that supports auto-initialization) +ResolutionErrors9.dfy(85,6): Error: when allocating an object of type 'ClassWithConstructor', one of its constructor methods must be called +ResolutionErrors9.dfy(86,6): Error: when allocating an object of type 'ClassWithConstructor', one of its constructor methods must be called +ResolutionErrors9.dfy(99,8): Error: when allocating an object of type 'Y', one of its constructor methods must be called +ResolutionErrors9.dfy(104,8): Error: when allocating an object of type 'Luci', one of its constructor methods must be called +ResolutionErrors9.dfy(105,8): Error: when allocating an object of type 'Luci', one of its constructor methods must be called +ResolutionErrors9.dfy(129,44): Error: type of right argument to * (bool) must agree with the result type (int) +ResolutionErrors9.dfy(210,16): Error: array selection requires integer- or bitvector-based numeric indices (got bool for index 0) +ResolutionErrors9.dfy(208,13): Error: incorrect type for selection into array (got real) +ResolutionErrors9.dfy(209,11): Error: incorrect type for selection into seq (got real) +ResolutionErrors9.dfy(210,23): Error: array selection requires integer- or bitvector-based numeric indices (got real for index 1) +ResolutionErrors9.dfy(212,17): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(212,17): Error: incorrect type for selection into array (got real) +ResolutionErrors9.dfy(212,22): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(212,22): Error: incorrect type for selection into array (got real) +ResolutionErrors9.dfy(213,17): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(213,17): Error: incorrect type for selection into array (got real) +ResolutionErrors9.dfy(214,19): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(214,19): Error: incorrect type for selection into array (got real) +ResolutionErrors9.dfy(215,15): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(215,15): Error: incorrect type for selection into seq (got real) +ResolutionErrors9.dfy(215,20): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(215,20): Error: incorrect type for selection into seq (got real) +ResolutionErrors9.dfy(216,15): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(216,15): Error: incorrect type for selection into seq (got real) +ResolutionErrors9.dfy(217,17): Error: wrong number of indices for multi-selection +ResolutionErrors9.dfy(217,17): Error: incorrect type for selection into seq (got real) +58 resolution/type errors detected in ResolutionErrors9.dfy