Trouble with Int Theory in Isabelle/HOL - isabelle

I am working on a proof which I was able to reduce to “of_int i = 0 ==> i = 0”. This seemed like a simple application of the rule “of_int_eq_0_iff” however I was unable to successfully apply this rule.
On further probing I found that I was unable to prove the following lemma
lemma of_int_eq_0_imp1: “of_int i = 0 ==> i = 0”
by any means whatsoever. That is, unless I declare the lemma within the context ring_char_0. Then the lemma can easily be proved as follows:
context ring_char_0 begin
lemma of_int_eq_0_imp1: “of_int i = 0 ==> i = 0”
using of_int_eq_iff [of i 0] by simp
end
But then I cannot apply this lemma outside of this context, which is what my main theorem requires (it resides within a different context).
Any help would be greatly appreciated.

The fact that you can only prove your lemma inside ring_char_0 should make you suspicious. The reason for this is that the lemma of_int_eq_0_iff is defined in the context of ring_char_0 itself. You can see this by typing e.g.
declare [[show_sorts]]
thm of_int_eq_0_iff
> (of_int (?z∷int) = (0∷?'a∷ring_char_0)) = (?z = (0∷int))
The reason for this is that in a ring with a characteristic k ≠ 0, this does not hold. In such a ring, of_int n will be zero for all multiples n of k, despite n not being 0.
If your original goal reduces to of_int i = 0 ==> i = 0 then maybe your original goal only holds for rings of characteristic 0, or you need a different proof, one that does not require of_int i = 0 ==> i = 0.

Related

VST Verification of Global Array of Doubles

I am currently attempting to use VST to verify the correctness of a project which involves a global array of doubles. However, when attempting to access the array I have that the head of the array is given as a data_at statement while the rest of the array is given as a sepcon list of mapsto statements and there does not appear to be any way to prove field_compatible for elements beyond the head of the array.
Trying to access elements beyond offset_val 0 seems to inevitably involve proving a size_compatible statement. This is where I run into a problem. Since the alignment of tdouble is set to 4 and the size is set to 8, there seems to be a possibility that the head of the array is at Ptrofs.modulus - 12 making size_compatible false for the next element in the array. Am I going about this the wrong way?
I made a toy example with the same problem that I've mentioned above.
double dbls[] = {0.0, 1.1};
int main() {
double sum;
sum = dbls[0] + dbls[1];
return 0;
}
I will frame my answer in the form of a Coq development:
Require Import VST.floyd.proofauto.
Require Import VST.progs.foo.
#[export] Instance CompSpecs : compspecs. make_compspecs prog. Defined.
Definition Vprog : varspecs. mk_varspecs prog. Defined.
Definition main_spec :=
DECLARE _main
WITH gv : globals
PRE [] main_pre prog tt gv
POST [ tint ] main_post prog gv.
Definition Gprog : funspecs := [ ].
Lemma body_main: semax_body Vprog Gprog f_main main_spec.
Proof.
start_function.
(* Remark 1: it seems to be a bug in VST 2.11.1 (and earlier versions)
that the array is not packaged up into
(data_at Ews (tarray Tdouble 2) ...)
the way it ought to be. This seems to work better for integer
arrays, et cetera.*)
(* Remark 2: you are right to be concerned about alignment, but
VST addresses that issue correctly. Any extern global variable
in a C program, such as your [dbls] array, is aligned at the
biggest possible alignment requirement. VST expresses this
with the "headptr" predicate, and for any identifier id,
(gv id) is a headptr. So therefore, *)
assert_PROP (headptr (gv _dbls)) by entailer!.
(* and you can see above the line, H: headptr (gv _dbls). *)
Print headptr.
(* This shows that (gv _dbls) must be at offset zero within
some block, which guarantees alignment at any type.
One useful consequence is proved by the lemma
headptr_field_compatible: *)
Check headptr_field_compatible.
(* And now, let's apply that lemma: *)
pose proof headptr_field_compatible (tarray tdouble 2) nil _
H (eq_refl _) Logic.I ltac:(simpl; rep_lia).
(* So we see that as long as the pesky 'align_compatible_rec' is proved,
the pointer (gv _dbls) should be 'field_compatible' with the array
type that you want. And it's straightforward though tedious to prove
the 'align_compatible_rec' premise, as follows: *)
spec H0.
apply align_compatible_rec_Tarray; intros.
Search align_compatible_rec.
eapply align_compatible_rec_by_value; [ reflexivity | ].
apply Z.divide_add_r.
apply Z.divide_0_r.
apply Z.divide_mul_l.
apply Z.mod_divide; compute; intros; congruence.
(* Normally, VST users shouldn't have to do this 'by hand'.
We should fix the bug (failure to nicely package the precondition).
But in the interim, perhaps this gives what you need for a workaround.*)

Problem with multi-sort atoms in Nominal Isabelle

I'm experimenting with Nominal2's support for multi-sort atoms (see this example in the Nominal package and the discussion in Section 5 of this paper). The idea is to define a toy typed λ-calculus in Church's style, that is, where each variable includes its type (e.g. λxₒ. xₒ). So, I came up with the following code:
theory Scratch
imports "Nominal2.Nominal2" "Nominal2.Atoms"
begin
(* In this example I'm using "var", "ty" and "Var" from Nominal2.Atoms *)
nominal_datatype exp =
EVar var
| EApp exp exp
| EAbs x::var M::exp binds x in M
(* I want to prove that λxₒ. xₒ = λyₒ. yₒ *)
lemma "EAbs (Var x (TVar ''o'')) (EVar (Var x (TVar ''o''))) = EAbs (Var y (TVar ''o'')) (EVar (Var y (TVar ''o'')))"
sorry (* cannot prove this! *)
end
As shown in the code above, the problem I found is that I cannot even prove the equality of two α-equivalent terms, which, I think, should work out of the box as is the case for single-sort atoms. Having a look at the Nominal2's implementation, it seems to me that the root cause is the lacking of a simproc similar to alpha_lst that works on at_base instead of at. However, since I'm not an expert in the Nominal package, I may very well be missing something.

How to make Pre and Post conditions for recursive functions in SPARK?

I'm translating an exercise I made in Dafny into SPARK, where one verifies a tail recursive function against a recursive one. The Dafny source (censored, because it might still be used for classes):
function Sum(n:nat):nat
decreases n
{
if n==0 then n else n+Sum(n-1)
}
method ComputeSum(n:nat) returns (s:nat)
ensures s == Sum(n)
{
s := 0;
// ...censored...
}
What I got in SPARK so far:
function Sum (n : in Natural) return Natural
is
begin
if n = 0 then
return n;
else
return n + Sum(n - 1);
end if;
end Sum;
function ComputeSum(n : in Natural) return Natural
with
Post => ComputeSum'Result = Sum(n)
is
s : Natural := 0;
begin
-- ...censored...
return s;
end ComputeSum;
I cannot seem to figure out how to express the decreases n condition (which now that I think about it might be a little odd... but I got graded for it a few years back so who am I to judge, and the question remains how to get it done). As a result I get warnings of possible overflow and/or infinite recursion.
I'm guessing there is a pre or post condition to be added. Tried Pre => n <= 1 which obviously does not overflow, but I still get the warning. Adding Post => Sum'Result <= n**n on top of that makes the warning go away, but that condition gets a "postcondition might fail" warning, which isn't right, but guess the prover can't tell. Also not really the expression I should check against, but I cannot seem to figure what other Post I'm looking for. Possibly something very close to the recursive expression, but none of my attempts work. Must be missing out on some language construct...
So, how could I express the recursive constraints?
Edit 1:
Following links to this SO answer and this SPARK doc section, I tried this:
function Sum (n : in Natural) return Natural
is
(if n = 0 then 0 else n + Sum(n - 1))
with
Pre => (n in 0 .. 2),
Contract_Cases => (n = 0 => Sum'Result = 0,
n >= 1 => Sum'Result = n + Sum(n - 1)),
Subprogram_Variant => (Decreases => n);
However getting these warnings from SPARK:
spark.adb:32:30: medium: overflow check might fail [reason for check: result of addition must fit in a 32-bits machine integer][#0]
spark.adb:36:56: warning: call to "Sum" within its postcondition will lead to infinite recursion
If you want to prove that the result of some tail-recursive summation function equals the result of a given recursive summation function for some value N, then it should, in principle, suffice to only define the recursive function (as an expression function) without any post-condition. You then only need to mention the recursive (expression) function in the post-condition of the tail-recursive function (note that there was no post-condition (ensures) on the recursive function in Dafny either).
However, as one of SPARK's primary goal is to proof the absence of runtime errors, you must have to prove that overflow cannot occur and for this reason, you do need a post-condition on the recursive function. A reasonable choice for such a post-condition is, as #Jeffrey Carter already suggested in the comments, the explicit summation formula for arithmetic progression:
Sum (N) = N * (1 + N) / 2
The choice is actually very attractive as with this formula we can now also functionally validate the recursive function itself against a well-known mathematically explicit expression for computing the sum of a series of natural numbers.
Unfortunately, using this formula as-is will only bring you somewhere half-way. In SPARK (and Ada as well), pre- and post-conditions are optionally executable (see also RM 11.4.2 and section 5.11.1 in the SPARK Reference Guide) and must therefore themselves be free of any runtime errors. Therefore, using the formula as-is will only allow you to prove that no overflow occurs for any positive number up until
max N s.t. N * (1 + N) <= Integer'Last <-> N = 46340
as in the post-condition, the multiplication is not allowed to overflow either (note that Natural'Last = Integer'Last = 2**31 - 1).
To work around this, you'll need to make use of the big integers package that has been introduced in the Ada 202x standard library (see also RM A.5.6; this package is already included in GNAT CE 2021 and GNAT FSF 11.2). Big integers are unbounded and computations with these integers never overflow. Using these integers, one can proof that overflow will not occur for any positive number up until
max N s.t. N * (1 + N) / 2 <= Natural'Last <-> N = 65535 = 2**16 - 1
The usage of these integers in a post-condition is illustrated in the example below.
Some final notes:
The Subprogram_Variant aspect is only needed to prove that a recursive subprogram will eventually terminate. Such a proof of termination must be requested explicitly by adding an annotation to the function (also shown in the example below and as discussed in the SPARK documentation pointed out by #egilhh in the comments). The Subprogram_Variant aspect is, however, not needed for your initial purpose: proving that the result of some tail-recursive summation function equals the result of a given recursive summation function for some value N.
To compile a program that uses functions from the new Ada 202x standard library, use compiler option -gnat2020.
While I use a subtype to constrain the range of permissible values for N, you could also use a precondition. This should not make any difference. However, in SPARK (and Ada as well), it is in general considered to be a best practise to express contraints using (sub)types as much as possible.
Consider counterexamples as possible clues rather than facts. They may or may not make sense. Counterexamples are optionally generated by some solvers and may not make sense. See also the section 7.2.6 in the SPARK user’s guide.
main.adb
with Ada.Numerics.Big_Numbers.Big_Integers;
procedure Main with SPARK_Mode is
package BI renames Ada.Numerics.Big_Numbers.Big_Integers;
use type BI.Valid_Big_Integer;
-- Conversion functions.
function To_Big (Arg : Integer) return BI.Valid_Big_Integer renames BI.To_Big_Integer;
function To_Int (Arg : BI.Valid_Big_Integer) return Integer renames BI.To_Integer;
subtype Domain is Natural range 0 .. 2**16 - 1;
function Sum (N : Domain) return Natural is
(if N = 0 then 0 else N + Sum (N - 1))
with
Post => Sum'Result = To_Int (To_Big (N) * (1 + To_Big (N)) / 2),
Subprogram_Variant => (Decreases => N);
-- Request a proof that Sum will terminate for all possible values of N.
pragma Annotate (GNATprove, Terminating, Sum);
begin
null;
end Main;
output (gnatprove)
$ gnatprove -Pdefault.gpr --output=oneline --report=all --level=1 --prover=z3
Phase 1 of 2: generation of Global contracts ...
Phase 2 of 2: flow analysis and proof ...
main.adb:13:13: info: subprogram "Sum" will terminate, terminating annotation has been proved
main.adb:14:30: info: overflow check proved
main.adb:14:32: info: subprogram variant proved
main.adb:14:39: info: range check proved
main.adb:16:18: info: postcondition proved
main.adb:16:31: info: range check proved
main.adb:16:53: info: predicate check proved
main.adb:16:69: info: division check proved
main.adb:16:71: info: predicate check proved
Summary logged in [...]/gnatprove.out
ADDENDUM (in response to comment)
So you can add the post condition as a recursive function, but that does not help in proving the absence of overflow; you will still have to provide some upper bound on the function result in order to convince the prover that the expression N + Sum (N - 1) will not cause an overflow.
To check the absence of overflow during the addition, the prover will consider all possible values that Sum might return according to it's specification and see if at least one of those value might cause the addition to overflow. In the absence of an explicit bound in the post condition, Sum might, according to its return type, return any value in the range Natural'Range. That range includes Natural'Last and that value will definitely cause an overflow. Therefore, the prover will report that the addition might overflow. The fact that Sum never returns that value given its allowable input values is irrelevant here (that's why it reports might). Hence, a more precise upper bound on the return value is required.
If an exact upper bound is not available, then you'll typically fallback onto a more conservative bound like, in this case, N * N (or use saturation math as shown in the Fibonacci example from the SPARK user manual, section 5.2.7, but that approach does change your function which might not be desirable).
Here's an alternative example:
example.ads
package Example with SPARK_Mode is
subtype Domain is Natural range 0 .. 2**15;
function Sum (N : Domain) return Natural
with Post =>
Sum'Result = (if N = 0 then 0 else N + Sum (N - 1)) and
Sum'Result <= N * N; -- conservative upper bound if the closed form
-- solution to the recursive function would
-- not exist.
end Example;
example.adb
package body Example with SPARK_Mode is
function Sum (N : Domain) return Natural is
begin
if N = 0 then
return N;
else
return N + Sum (N - 1);
end if;
end Sum;
end Example;
output (gnatprove)
$ gnatprove -Pdefault.gpr --output=oneline --report=all
Phase 1 of 2: generation of Global contracts ...
Phase 2 of 2: flow analysis and proof ...
example.adb:8:19: info: overflow check proved
example.adb:8:28: info: range check proved
example.ads:7:08: info: postcondition proved
example.ads:7:45: info: overflow check proved
example.ads:7:54: info: range check proved
Summary logged in [...]/gnatprove.out
I landed in something that sometimes works, which I think is enough for closing the title question:
function Sum (n : in Natural) return Natural
is
(if n = 0 then 0 else n + Sum(n - 1))
with
Pre => (n in 0 .. 10), -- works with --prover=z3, not Default (CVC4)
-- Pre => (n in 0 .. 100), -- not working - "overflow check might fail, e.g. when n = 2"
Subprogram_Variant => (Decreases => n),
Post => ((n = 0 and then Sum'Result = 0)
or (n > 0 and then Sum'Result = n + Sum(n - 1)));
-- Contract_Cases => (n = 0 => Sum'Result = 0,
-- n > 0 => Sum'Result = n + Sum(n - 1)); -- warning: call to "Sum" within its postcondition will lead to infinite recursion
-- Contract_Cases => (n = 0 => Sum'Result = 0,
-- n > 0 => n + Sum(n - 1) = Sum'Result); -- works
-- Contract_Cases => (n = 0 => Sum'Result = 0,
-- n > 0 => Sum'Result = n * (n + 1) / 2); -- works and gives good overflow counterexamples for high n, but isn't really recursive
Command line invocation in GNAT Studio (Ctrl+Alt+F), --counterproof=on and --prover=z3 my additions to it:
gnatprove -P%PP -j0 %X --output=oneline --ide-progress-bar --level=0 -u %fp --counterexamples=on --prover=z3
Takeaways:
Subprogram_Variant => (Decreases => n) is required to tell the prover n decreases for each recursive invocation, just like the Dafny version.
Works inconsistently for similar contracts, see commented Contract_Cases.
Default prover (CVC4) fails, using Z3 succeeds.
Counterproof on fail makes no sense.
n = 2 presented as counterproof for range 0 .. 100, but not for 0 .. 10.
Possibly related to this mention in the SPARK user guide: However, note that since the counterexample is always generated only using CVC4 prover, it can just explain why this prover cannot prove the property.
Cleaning between changing options required, e.g. --prover.

How to handle "Forall (closed_wrt_vars (eq _z')) P " goals in VST?

This time I'm proving function calling other. vars.c:
int pure0 ()
{
return 0;
}
int get0(int* arr)
{
int z = pure0();
return 0;
}
My proof start - verif_vars.v:
Require Import floyd.proofauto.
Require Import vars.
Local Open Scope logic.
Local Open Scope Z.
Definition get0_spec :=
DECLARE _get0
WITH sh : share, arr : Z->val, varr : val
PRE [_arr OF (tptr tint)]
PROP ()
LOCAL (`(eq varr) (eval_id _arr);
`isptr (eval_id _arr))
SEP (`(array_at tint sh arr 0 100) (eval_id _arr))
POST [tint] `(array_at tint sh arr 0 100 varr).
Definition pure0_spec :=
DECLARE _pure0
WITH sh : share
PRE []
PROP ()
LOCAL ()
SEP ()
POST [tint] local(`(eq (Vint (Int.repr 0))) retval).
Definition Vprog : varspecs := nil.
Definition Gprog : funspecs := get0_spec :: pure0_spec ::nil.
Lemma body_pure0: semax_body Vprog Gprog f_pure0 pure0_spec.
Proof.
start_function.
forward.
Qed.
Lemma body_get0: semax_body Vprog Gprog f_get0 get0_spec.
Proof.
start_function.
name arrarg _arr.
forward_call (sh).
entailer!.
Which induces the goal:
2 subgoals, subgoal 1 (ID 566)
Espec : OracleKind
sh : share
arr : Z -> val
varr : val
Delta := abbreviate : tycontext
POSTCONDITION := abbreviate : ret_assert
MORE_COMMANDS := abbreviate : statement
Struct_env := abbreviate : type_id_env.type_id_env
arrarg : name _arr
============================
Forall (closed_wrt_vars (eq _z')) [`(array_at tint sh arr 0 100 varr)]
subgoal 2 (ID 567) is:
DO_THE_after_call_TACTIC_NOW
I suppose it states, that the function call does not alter arr contents, which is quite obvious for me.
What can I do with this goal? Which tactic applies here, and what exactly means the statement? Should I enrich the pure0 spec to somehow point out, that it does not modify anything?
FIRST: When writing VST/Verifiable-C questions, please indicate which version of VST you are using. It appears you are using 1.4.
SECOND: I am not sure this answers all your questions, but,
"closed_wrt_vars S P" says that the lifted assertion P is closed with respect to all the variables in the set S. That is, S is a set of C-language identifiers that may stand for nonaddressable local variables ("temps", not "vars"). P is an assertion of the form "environ->mpred", and "closed" means that if you change the "environ" to have different values for any of the variables in set S, then the truth of P will not change.
"Forall" is Coq's standard library predicate to apply a predicate to a list. So,
Forall (closed_wrt_vars (eq _z')) [`(array_at tint sh arr 0 100 varr)]
means, let the set S be the singleton set containing just the variable _z'.
We assert here that all the predicates in the list are closed w.r.t. S.
There's exactly one predicate in the list, and it's "trivially lifted",
that is, for any predicate (P: mpred), the lifted predicate
`(P)
is equivalent to (fun rho:environ => P). Trivially, then, `P doesn't
care what you do to rho, including changing the value of _z'.
The "auto with closed" (or just to be sure, "auto 50 with closed")
should take care of this, and you indicate that it does take care of it.
So I assume that the rest of your question was, "what's going on here?",
and I hope I answered it.
Solution, used in vst/progs/verif_reverse.v:
auto with closed.
Unfortunately, it answers only half of the questions.
By the way (unrelated to your question), the precondition
`isptr (eval_id _arr) for get0 is probably unnecessary.
It is implied already by `(array_at tint sh arr 0 100) (eval_id _arr)).
Furthermore, suppose you did want the `isptr (eval_id _arr) in your precondition; you might consider writing it as,
PROP (isptr varr)
LOCAL (`(eq varr) (eval_id _arr))
SEP (`(array_at tint sh arr 0 100 varr))
which is (in some ways) simpler and more "canonical".

Trying to understand fix/assume/show "Failure to refine goal"; Cmd to show proof info for schematic vars

My questions here are a spin-off of what was a tangent in my previous question.
For these questions, I use a very simple lemma, though my second question is fairly involved.
The error "Local statement fails to refine any pending goal" is one of the more frustrating errors, when using show or thus and when all my logic appears correct, so I'm trying to better understand the error message that occurs below at lemma fix_2.
I know the fix for it, which is my lemma fix_1, but the more I know, the better I might be at dealing with another "fails to refine any pending goal".
Anyone interested might be able to answer Q1 and Q2 just by reading the questions, and then looking at the two lemmas.
There's a lot of information I put in below. I format the comments below like I do to be able to compare how this and goal change after commands. Doing that, I'm able to get a better understanding of what's happening with the use of let, def, and fix/assume, where the main purpose is to try and understanding the error at the show command of lemma fix_2.
I don't know how to make a question like this simple.
Questions
Here, I put forth two questions. You will need to skip down to look at lemma let_1 and lemma fix_2. I tried to use HTML anchors to create links within this page, but it didn't work.
Q1: Below, at lemma let_1, I use print_commands. I looked through those commands to try to find commands that will give me information about how schematic variables are instantiated. I found print_binds, which shows my use of schematic variable ?w. Are there any other commands that will give me information about what's happening with schematic variables in a proof?
Q2: Am I right in saying the following?
At the use of the show "card {} = 0" in lemma fix_2, the this fact and its implicit hypothesis, card w = 0 [w == {}], are used to create a rule similar to what is exported after the {...} block in lemma fix_1, where the exported rule there is (?w2 == {}) ==> card ?w2 = 0.
The created rule is then used to do some kind of unification with the proof goal in show "card {} = 0", in which the schematic variable is instantiated with {}, but something doesn't match up, and an error occurs.
The behavior of def and fix/assume described
The primary context of this question is this statement by L.Noschinski on the IsaUserList:
When you use "fix" or "def" to define a variable, they either get just generalized (i.e. turned into schematics) (fix) or replaced by their right hand side (definitions) when a block is closed / a show is performed.
I partially restate it to show how I understand it for the show command, where what I say is also based on how def_1 and fix_1 behave below, which both use a {...} block:
If the statements def x == "P" and fix y assume "y == P" are used before a show command, as in def_2 and fix_2 below, at the use of the show command, the following will happen:
For def x, any use of x in the this fact will be replaced by P.
For fix y, a rule will be created using both the this fact and its implicit hypothesis, such as after the {...} in fix_1 below. In this rule, y will be replaced by a schematic variable.
Five lemmas for studying this, goal, have and show
declares
I use show_question_marks because I need to see when schematic variables are introduced when fix/assume is used.
The use of show_hyps shows the implicit hypothesis for a proof fact, in square brackets. These implicit conditions get used in an exported rule when fix/assume is used.
declare[[show_question_marks=true, show_hyps=true]]
declare[[show_sorts=false, show_types=false, show_brackets=false]]
One basic lemma, and 5 variations which have a proof fact not used
The basic lemma shows what's actually being proved. I then have the following:
let_1 to look at how the use of let affects this.
def_1, which uses def and a {...} block, to see how def behaves and how this is exported.
def_2, which has no {...} block, to be able to look at this before show is used.
fix_1 and fix_2, which are likewise to def_1 and def_2, but use fix/assume.
lemma with no proof fact
All of the lemmas below are the next lemma, and are proved the same. What the others have in addition is one proof fact which is not needed, and not used by the show command.
The purpose of the proof fact is to help me see how def and fix change the this fact when have is proved, and see how they export this after a block {...} is closed.
lemma "card {} = (0::nat)"
proof-
show "card {} = 0"
by(simp)
qed
let_1
lemma let_1: "card {} = (0::nat)"
proof-
let ?w = "{}::'a set" (*No `this` fact: ?w is instantiated as {}.*)
print_commands
print_binds (*term bindings: w? == bot *)
have "card ?w = (0::nat)" (*goal: card {} = 0 *)
by(simp) (*this: card {} = 0 *)
show "card {} = 0" (*goal: card {} = 0 *)
by(simp)
qed
def_1, {...}
lemma def_1: "card {} = (0::nat)"
proof-
{def w == "{}::'a set" (*this: w == {} [w == {}] [name "local.w_def"] *)
from this
have "card w = (0::nat)" (*goal: card w = 0 *)
by(simp) (*this: card w = 0 [w == {}] *)
} (*this: card {} = 0 *)
show "card {} = 0" (*goal: card {} = 0 *)
by(simp)
qed
def_2, no block
lemma def_2: "card {} = (0::nat)"
proof-
def w == "{}::'a set" (*this: w == {} [w == {}] [name "local.w_def"] *)
from this
have "card w = (0::nat)" (*goal: card w = 0 *)
by(simp) (*this: card w = 0 [w == {}] *)
show "card {} = 0" (*goal: card {} = 0 *)
by(simp)
qed
fix_1, {...}
lemma fix_1: "card {} = (0::nat)"
proof-
{fix w assume "w == {}::'a set" (*this: w == {} [w == {}] *)
from this
have "card w = (0::nat)" (*goal: card w = 0 *)
by(simp) (*this: card w = 0 [w == {}] *)
} (*this: (?w2 == {}) ==> card ?w2 = 0 *)
show "card {} = 0" (*goal: card {} = 0 *)
by(simp)
qed
fix_2, no block
lemma fix_2: "card {} = (0::nat)"
proof-
fix w assume "w == {}::'a set"(*this: w == {} [w == {}] *)
from this
have "card w = (0::nat)" (*goal: card w = 0 *)
by(simp) (*this: card w = 0 [w == {}] *)
show "card {} = 0" (*Local statement fails to refine any pending goal
Failed attempt to solve goal by exported rule:
(?w3 == {}) ==> card {} = 0 *)
oops
Filling in details for L.Paulson's answer, per my understanding
140119:
The answer to Q2 and why the error occurs in fix_2 is given by L.Paulson when he says,
In the proof of fix_2, you have "fix w". This extends the context with w. The context no longer matches the original context of the goal...
After having done a search on "local context" in isar-ref.pdf, I do the gofer work to fill in some details, as I understand them.
The explicit answer to my Q2 is, no, I'm not right, where I'll quote from isar-ref.pdf to explain why the formula (?w3 == {}) ==> card {} = 0 is in the error message.
Another short answer from isar-ref.pdf is that presume in place of assume will "weaken the local context" and get rid of the error, because, apparently, the context is no longer extended by w, as described by L.Paulson.
Why the complexity of my example fix_2
My setup for this question was academic, as if a professor said, "In fix_2, modify the lemma in a minimal way to get rid of the error, while still using fix. In particular, do not use def, obtain, or let to eliminate the error. My use of {...}, as in fix_1, was an acceptable solution, but I wanted to go further and understand what produces the formula in the error message of fix_2, to help me in the future.
My use of fix/assume in fix_2 is specific to the previous question I link to at the top. Here, I introduced a proof fact to ensure an error, like I get there, but here, I don't use the proof fact, to simplify things, so I don't have to use from this or thus, but only have to use show.
In my answer to my previous question, for the lemma at hand, I couldn't see why def produced no error when fix/assume did. The info in the output panel is almost identical there for def and fix/assume, and isar-ref.pdf describes def as an abbreviation for fix/assume, page 117, where the key word is "basically":
Basically, def x == t abbreviates fix x "assume x == t"...
Local context, it' like, not unimportant, which is to say huge
There was this foggy notion in my mind that somewhere in the answer was the issue of local context, because I've seen the word "context" a lot in the docs. Knowing that is why I threw in the use of {...} to fix the problem once I read L.Noschinski's tip, and why I searched on "local context" after L.Paulson's answer.
Why (?w3 == {}) ==> card {} = 0 is in the error message, I think
The explanation is on page 34 of isar-ref.pdf in 2.2.2 Structured statements:
A simple statement consists of named propositions. The full form admits local context elements followed by the actual conclusions, such as fixes x assumes "A x" shows "B x". The final result emerges as a Pure rule after discharging the context: !!x. A x ==> B x.
I'm assuming here that fixe/assume in a proof works similar to a fixes/assumes in a lemma statement.
My first guess, rather than my Q2, was on track. By L.Noschinski, I knew that w in assume "w == {}" gets changed to a schematic variable when show is invoked, so the left-hand side of the error formula, (?w3 == {}) ==> card {} = 0, matches up with with the formula of assume "w == {}", and the right-hand side matches up with the formula of show "card {} = 0".
I changed my guess on a bicycle ride, that a rule like fix_1's exported this was created, (?w2 == {}) ==> card ?w2 = 0, and that ?w2 was instantiated with {}. That didn't match very well with (?w3 == {}) ==> card {} = 0, but it was all a guess anyway.
Two more quotes about local context
Because of my use of fix/assume, this next quote put me on the right track, after given the answer by L.Paulson that my fix/assume extends the context.
pg.32 in 2.2 The Isar proof language:
The remaining elements fix and assume build up a local context (see §2.2.1),
while show refines a pending sub-goal by the rule resulting from a nested
sub-proof (see §2.2.3).
My pursuit here has largely been about trying to see what's happening under the hood when show is used. I'm still pursuing commands to give me some feedback about show that's not normally shown in the output panel:
pg.117 in 6.2.4 Goals:
show a: alpha is like have a: alpha plus a second stage to refine some
pending sub-goal for each one of the finished result...
Using presume in place of assume
Finally, I show another solution to my academic exercise. What I quote below wouldn't have had any meaning to me without L.Paulson's answer that in fix_2, I'm extending the context.
pg.120 in 6.3.2 Initial and terminal proof steps:
Debugging such a situation might involve temporarily changing show into have,
or weakening the local context by replacing occurrences of assume by presume.
lemma fix_3: "card {} = (0::nat)"
proof-
fix w presume "w == {}::'a set"(*this: w == {} [w == {}] *)
from this
have "card w = (0::nat)" (*goal: card w = 0 *)
by(simp) (*this: card w = 0 [w == {}] *)
show "card {} = 0" (*goal: card {} = 0 *)
print_binds (*term bindings:
?this == card w = 0
... == 0
?thesis == card bot = 0 *)
by(simp)
You are making something that should be simple very complicated.
In the proof of fix_2, you have fix w. This extends the context with w. The context no longer matches the original context of the goal, hence show card {} = 0 yields the error message you mention.
Probably the lemma you really need in this proof is !!w. card w = 0. When proving this lemma, you will insert fix w in a local block, and it will all work.

Resources