Working with Isabelle's code generator: Data refinement and higher order functions - code-generation

This is a follow-up on Isabelle's Code generation: Abstraction lemmas for containers?:
I want to generate code for the_question in the following theory:
theory Scratch imports Main begin
typedef small = "{x::nat. x < 10}" morphisms to_nat small
by (rule exI[where x = 0], simp)
code_datatype small
lemma [code abstype]: "small (to_nat x) = x" by (rule to_nat_inverse)
definition a_pred :: "small ⇒ bool"
where "a_pred = undefined"
definition "smaller j = [small i . i <- [0 ..< to_nat j]]"
definition "the_question j = (∀i ∈ set (smaller j). a_pred j)"
The problem is that the equation for smaller is not suitable for code generation, as it mentions the abstraction function small.
Now according to Andreas’ answer to my last question and the paper on data refinement, the next step is to introduce a type for sets of small numbers, and create a definition for smaller in that type:
typedef small_list = "{l. ∀x∈ set l. (x::nat) < 10}" by (rule exI[where x = "[]"], auto)
code_datatype Abs_small_list
lemma [code abstype]: "Abs_small_list (Rep_small_list x) = x" by (rule Rep_small_list_inverse)
definition "smaller' j = Abs_small_list [ i . i <- [0 ..< to_nat j]]"
lemma smaller'_code[code abstract]: "Rep_small_list (smaller' j) = [ i . i <- [0 ..< to_nat j]]"
unfolding smaller'_def
by (rule Abs_small_list_inverse, cases j, auto elim: less_trans simp add: small_inverse)
Now smaller' is executable. From what I understand I need to redefine operations on small list as operations on small_list:
definition "small_list_all P l = list_all P (map small (Rep_small_list l))"
lemma[code]: "the_question j = small_list_all a_pred (smaller' j)"
unfolding small_list_all_def the_question_def smaller'_code smaller_def Ball_set by simp
I can define a good looking code equation for the_question. But the definition of small_list_all is not suitable for code generation, as it mentions the abstraction morphismsmall. How do I make small_list_all executable?
(Note that I cannot unfold the code equation of a_pred, as the problem actually occurs in the code equation of the actually recursive a_pred. Also, I’d like to avoid hacks that involve re-checking the invariant at runtime.)

I don't have a good solution to the general problem, but here's an idea that will let you generate code for the_question in this particular case.
First, define a function predecessor :: "small ⇒ small with an abstract code equation (possibly using lift_definition from λn::nat. n - 1).
Now you can prove a new code equation for smaller whose rhs uses if-then-else, predecessor and normal list operations:
lemma smaller_code [code]:
"smaller j = (if to_nat j = 0 then []
else let k = predecessor j in smaller k # [k])"
(More efficient implementations are of course possible if you're willing to define an auxiliary function.)
Code generation should now work for smaller, since this code equation doesn't use function small.

The short answer is no, it does not work.
The long answer is that there are often workarounds possible. One is shown by Brian in his answer. The general idea seems to be
Separate the function that has the abstract type in covariant positions besides the final return value (i.e. higher order functions or functions returning containers of abstract values) into multiple helper functions so that abstract values are only constructed as a single return value of one of the helper function.
In Brian’s example, this function is predecessor. Or, as another simple example, assume a function
definition smallPrime :: "nat ⇒ small option"
where "smallPrime n = (if n ∈ {2,3,5,7} then Some (small n) else None)"
This definition is not a valid code equation, due to the occurrence of small. But this derives one:
definition smallPrimeHelper :: "nat ⇒ small"
where "smallPrimeHelper n = (if n ∈ {2,3,5,7} then small n else small 0)"
lemma [code abstract]: "to_nat (smallPrimeHelper n) = (if n ∈ {2,3,5,7} then n else 0)"
by (auto simp add: smallPrimeHelper_def intro: small_inverse)
lemma [code_unfold]: "smallPrime n = (if n ∈ {2,3,5,7} then Some (smallPrimeHelper n) else None)"
unfolding smallPrime_def smallPrimeHelper_def by simp
If one wants to avoid the redundant calculation of the predicate (which might be more complex than just ∈ {2,3,5,7}, one can make the return type of the helper smarter by introducing an abstract view, i.e. a type that contains both the result of the computation, and the information needed to construct the abstract type from it:
typedef smallPrime_view = "{(x::nat, b::bool). x < 10 ∧ b = (x ∈ {2,3,5,7})}"
by (rule exI[where x = "(2, True)"], auto)
setup_lifting type_definition_small
setup_lifting type_definition_smallPrime_view
For the view we have a function building it and accessors that take the result apart, with some lemmas about them:
lift_definition smallPrimeHelper' :: "nat ⇒ smallPrime_view"
is "λ n. if n ∈ {2,3,5,7} then (n, True) else (0, False)" by simp
lift_definition smallPrimeView_pred :: "smallPrime_view ⇒ bool"
is "λ spv :: (nat × bool) . snd spv" by auto
lift_definition smallPrimeView_small :: "smallPrime_view ⇒ small"
is "λ spv :: (nat × bool) . fst spv" by auto
lemma [simp]: "smallPrimeView_pred (smallPrimeHelper' n) ⟷ (n ∈ {2,3,5,7})"
by transfer simp
lemma [simp]: "n ∈ {2,3,5,7} ⟹ to_nat (smallPrimeView_small (smallPrimeHelper' n)) = n"
by transfer auto
lemma [simp]: "n ∈ {2,3,5,7} ⟹ smallPrimeView_small (smallPrimeHelper' n) = small n"
by (auto intro: iffD1[OF to_nat_inject] simp add: small_inverse)
With that we can derive a code equation that does the check only once:
lemma [code]: "smallPrime n =
(let spv = smallPrimeHelper' n in
(if smallPrimeView_pred spv
then Some (smallPrimeView_small spv)
else None))"
by (auto simp add: smallPrime_def Let_def)

Related

How do I code algebraic cpos as an Isabelle locale

I am trying to prove the known fact that there is a standard way to build an algebriac_cpo from a partial_order. My problem is I keep getting the error
Type unification failed: No type arity set :: partial_order
and I can not understand what this means.
I think I have tracked down my problem to the definition of cpo. The definition works and I have proven various results for it but the working interpretation of a partial_order fails with cpo.
theory LocaleProblem imports "HOL-Lattice.Bounds"
begin
definition directed:: "'a::partial_order set ⇒ bool" where
" directed a ≡
¬a={} ∧ ( ∀ a1 a2. a1∈a∧ a2∈a ⟶ (∃ ub . (is_Sup {a1,a2} ub))) "
class cpo = partial_order +
fixes bot :: "'a::partial_order" ("⊥")
assumes bottom: " ∀ x::'a. ⊥ ⊑ x "
assumes dlub: "directed (d::'a::partial_order set) ⟶ (∃ lub . is_Inf d lub) "
print_locale cpo
interpretation "idl" : partial_order
"(⊆)::( ('b set) ⇒ ('b set) ⇒ bool) "
by (unfold_locales , auto) (* works *)
interpretation "idl" : cpo
"(⊆)::( ('b set) ⇒ ('b set) ⇒ bool) "
"{}" (* gives error
Type unification failed: No type arity set :: partial_order
Failed to meet type constraint:
Term: (⊆) :: 'b set ⇒ 'b set ⇒ bool
Type: ??'a ⇒ ??'a ⇒ bool *)
Any help much appreciated. david
You offered two solutions.
Following the work of Hennessy in Algebraic Theory of Processes" I am trying to prove (where I(A) are the Ideal which are sets) "If a is_a partial_order then I(A) is an algebraic_cpo" I will then want to apply the result to a number of semantics, give as sets. Dose your comment mean that the second solution is not a good idea?
Initially I had a proven lemma that started
lemma directed_ran: "directed (d::('a::partial_order×'b::partial_order) set) ⟹ directed (Range d)"
proof (unfold directed_def)
With the First solution started well:
context partial_order begin (* type 'a is now a partial_order *)
definition
is_Sup :: "'a set ⇒ 'a ⇒ bool" where
"is_Sup A sup = ((∀x ∈ A. x ⊑ sup) ∧ (∀z. (∀x ∈ A. x ⊑ z) ⟶ sup ⊑ z))"
definition directed:: "'a set ⇒ bool" where
" directed a ≡
¬a={} ∧ ( ∀ a1 a2. a1∈a∧ a2∈a ⟶ (∃ ub . (is_Sup {a1,a2} ub))) "
lemma directed_ran: "directed (d::('c::partial_order×'b::partial_order) set) ⟹ directed (Range d)"
proof - assume a:"directed d"
from a local.directed_def[of "d"] (* Fail with message below *)
show "directed (Range d)"
Alas the working lemma now fails: I rewrote
proof (unfold local.directed_def)
so I explored and found that although the fact local.directed_def exists is can not be unified
Failed to meet type constraint:
Term: d :: ('c × 'b) set
Type: 'a set
I changed the type successfully in the lemma statement but now can find no way to unfold the definition in the proof. Is there some way to do this?
Second solution again starts well:
instantiation set:: (type) partial_order
begin
definition setpoDef: "a⊑(b:: 'a set) = subset_eq a b"
instance proof
fix x::"'a set" show " x ⊑ x" by (auto simp: setpoDef)
fix x y z::"'a set" show "x ⊑ y ⟹ y ⊑ z ⟹ x ⊑ z" by(auto simp: setpoDef)
fix x y::"'a set" show "x ⊑ y ⟹ y ⊑ x ⟹ x = y" by (auto simp: setpoDef)
qed
end
but next step fails
instance proof
fix d show "directed d ⟶ (∃lub. is_Sup d (lub::'a set))"
proof assume "directed d " with directed_def[of "d"] show "∃lub. is_Sup d lub"
by (metis Sup_least Sup_upper is_SupI setpoDef)
qed
next
from class.cpo_axioms_def[of "bot"] show "∀x . ⊥ ⊑ x " (* Fails *)
qed
end
the first subgoal is proven but show "∀x . ⊥ ⊑ x " although cut an paste from the subgaol in the output does not match the subgoal. Normally at this point is need to add type constraints. But I cannot find any that work.
Do you know what is going wrong?
Do you know if I can force the Output to reveal the type information in the sugoal.
The interpretation command acts on the locale that a type class definition implicitly declares. In particular, it does not register the type constructor as an instance of the type class. That's what the command instantiation is good for. Therefore, in the second interpretation, Isabelle complains about set not having been registered as a instance of partial_order.
Since directed only needs the ordering for one type instance (namely 'a), I recommend to move the definition of directed into the locale context of the partial_order type class and remove the sort constraint on 'a:
context partial_order begin
definition directed:: "'a set ⇒ bool" where
"directed a ≡ ¬a={} ∧ ( ∀ a1 a2. a1∈a∧ a2∈a ⟶ (∃ ub . (is_Sup {a1,a2} ub))) "
end
(This only works if is_Sup is also defined in the locale context. If not, I recommend to replace the is_Sup condition with a1 <= ub and a2 <= ub.)
Then, you don't need to constrain 'a in the cpo type class definition, either:
class cpo = partial_order +
fixes bot :: "'a" ("⊥")
assumes bottom: " ∀ x::'a. ⊥ ⊑ x "
assumes dlub: "directed (d::'a set) ⟶ (∃ lub . is_Inf d lub)"
And consequently, your interpretation should not fail due to sort constraints.
Alternatively, you could declare set as an instance of partial_order instead of interpreting the type class. The advantage is that you can then also use constants and theorems that need partial_order as a sort constraint, i.e., that have not been defined or proven inside the locale context of partial_order. The disadvantage is that you'd have to define the type class operation inside the instantiation block. So you can't just say that the subset relation should be used; this has to be a new constant. Depending on what you intend to do with the instantiation, this might not matter or be very annoying.

Cardinality of finFieldType / Euler's criterion

I am trying to prove a very limited form of Euler's criterion:
Variable F : finFieldType.
Hypothesis HF : (1 != -1 :> F).
Lemma euler (a : F) : a^+(#|F|.-1./2) = -1 -> forall x, x^+2 != a.
I have the bulk of the proof done already, but I am left with odd (#|F|.-1) = 0, that is, #|F|.-1 is even. (I'm not interested in characteristic 2). I can't seem to find useful facts in the math comp library about the cardinality of finFieldTypes. For example, I would expect a lemma saying there exists a p such that prime p and #|F| = p. Am I missing something here?
By the way, I could also have totally missed an already existing proof of Euler's criterion in the math comp library itself.
I am not aware of a proof of Euler's criterion, but I found two lemmas, in finfield, that give you the two results that you expected in order to finish your proof (the second one may not be presented as you expected):
First, you have the following lemma which gives you the prime p corresponding to the characteristic of your field F (as long as it is a finFieldType):
Lemma finCharP : {p | prime p & p \in [char F]}.
Then, another lemma gives you the cardinality argument :
Let n := logn p #|R|.
Lemma card_primeChar : #|R| = (p ^ n)%N.
The problem with the second lemma is that your field should be recognized as a PrimeCharType, which roughly corresponds to a ringType with an explicit characteristic. But given the first lemma, you are able to give such a structure to your field (which canonically has a ringType), on the fly. A possible proof could be the following
Lemma odd_card : ~~ odd (#|F|.-1).
Proof.
suff : odd (#|F|) by have /ltnW/prednK {1}<- /= := finRing_gt1 F.
have [p prime_p char_F] := (finCharP F); set F_pC := PrimeCharType p_char.
have H : #|F| = #|F_primeChar| by []; rewrite H card_primeChar -H odd_exp => {H F_pC}.
apply/orP; right; have := HF; apply: contraR=> /(prime_oddPn prime_p) p_eq2.
by move: char_F; rewrite p_eq2=> /oppr_char2 ->.
Qed.

Computing with a finite subset of an infinite representation in Coq

I have a function Z -> Z -> whatever which I treat as a sort of a map from (Z, Z) to whatever, let's type it as FF.
With whatever being a simple sum constructible from nix or inj_whatever.
This map I initialize with some data, in the fashion of:
Definition i (x y : Z) (f : FF) : FF :=
fun x' y' =>
if andb (x =? x') (y =? y')
then inj_whatever
else f x y.
The =? represents boolean decidable equality on Z, from Coq's ZArith.
Now I would like to have equality on two of such FFs, I don't mind invoking functional_extensionality. What I would like to do now is to have Coq computationally decide equality of two FFs.
For example, suppose we do something along the lines of:
Definition empty : FF := fun x y => nix.
Now we add some arbitrary values to make foo and foo', those are equivalent under functional extensionality:
Definition foo := i 0 0 (i 0 (-42) (i 56 1 empty)).
Definition foo' := i 0 (-42) (i 56 1 (i 0 0 empty)).
What is a good way to automatically have Coq determine foo = foo'. Ltac level stuff? Actual terminating computation? Do I need domain restriction to a finite one?
The domain restriction is a bit of an intricate one. I manipulate the maps in a way f : FF -> FF, where f can extend the subset of Z x Z that the computation is defined on. As such, come to think of it, it can't be f : FF -> FF, but more like f : FF -> FF_1 where FF_1 is a subset of Z x Z that is extended by a small constant. As such, when one applies f n times, one ends up with FF_n which is equivalent to domain restriction of FF plus n * constant to the domain. So the function f slowly (by a constant factor) expands the domain FF is defined on.
As I said in the comment more specifics are needed in order to elaborate a satisfactory answer. See the below example --- intended for a step by step description --- on how to play with equality on restricted function ranges using mathcomp:
From mathcomp Require Import all_ssreflect all_algebra.
Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
(* We need this in order for the computation to work. *)
Section AllU.
Variable n : nat.
(* Bounded and unbounded fun *)
Definition FFb := {ffun 'I_n -> nat}.
Implicit Type (f : FFb).
Lemma FFP1 f1 f2 : reflect (f1 = f2) [forall x : 'I_n, f1 x == f2 x].
Proof. exact/(equivP eqfunP)/ffunP. Qed.
Lemma FFP2 f1 f2 :
[forall x : 'I_n, f1 x == f2 x] = all [fun x => f1 x == f2 x] (enum 'I_n).
Proof.
by apply/eqfunP/allP=> [eqf x he|eqf x]; apply/eqP/eqf; rewrite ?enumT.
Qed.
Definition f_inj (f : nat -> nat) : FFb := [ffun x => f (val x)].
Lemma FFP3 (f1 f2 : nat -> nat) :
all [fun x => f1 x == f2 x] (iota 0 n) -> f_inj f1 = f_inj f2.
Proof.
move/allP=> /= hb; apply/FFP1; rewrite FFP2; apply/allP=> x hx /=.
by rewrite !ffunE; apply/hb; rewrite mem_iota ?ltn_ord.
Qed.
(* Exercise, derive bounded eq from f_inj f1 = f_inj f2 *)
End AllU.
The final lemma should indeed allow you reduce equality of functions to a computational, fully runnable Gallina function.
A simpler version of the above, and likely more useful to you is:
Lemma FFP n (f1 f2 : nat -> nat) :
[forall x : 'I_n, f1 x == f2 x] = all [pred x | f1 x == f2 x] (iota 0 n).
Proof.
apply/eqfunP/allP=> eqf x; last by apply/eqP/eqf; rewrite mem_iota /=.
by rewrite mem_iota; case/andP=> ? hx; have /= -> := eqf (Ordinal hx).
Qed.
But it depends on how you (absent) condition on range restriction is specified.
After your edit, I think I should add a note on the more general topic of map equality, indeed you can define a more specific type of maps other than A -> B and then build a decision procedure.
Most typical map types [including the ones in the stdlib] will work, as long as they support the operation of "binding retrieval", so you can reduce equality to the check of finitely-many bound values.
In fact, the maps in Coq's standard library do already provide you such computational equality function.
Ok, this is a rather brutal solution which does not attempt to avoid doing the same case distinctions multiple times but it's fully automated.
We start with a tactic which inspects whether two integers are equal (using Z.eqb) and translates the results to a proposition which omega can deal with.
Ltac inspect_eq y x :=
let p := fresh "p" in
let q := fresh "q" in
let H := fresh "H" in
assert (p := proj1 (Z.eqb_eq x y));
assert (q := proj1 (Z.eqb_neq x y));
destruct (Z.eqb x y) eqn: H;
[apply (fun p => p eq_refl) in p; clear q|
apply (fun p => p eq_refl) in q; clear p].
We can then write a function which fires the first occurence of i it can find. This may introduce contradictory assumptions in the context e.g. if a previous match has revealed x = 0 but we now call inspect x 0, the second branch will have both x = 0 and x <> 0 in the context. It will be automatically dismissed by omega.
Ltac fire_i x y := match goal with
| [ |- context[i ?x' ?y' _ _] ] =>
unfold i at 1; inspect_eq x x'; inspect_eq y y'; (omega || simpl)
end.
We can then put everything together: call functional extensionality twice, repeat fire_i until there's nothing else to inspect and conclude by reflexivity (indeed all the branches with contradictions have been dismissed automatically!).
Ltac eqFF :=
let x := fresh "x" in
let y := fresh "y" in
intros;
apply functional_extensionality; intro x;
apply functional_extensionality; intro y;
repeat fire_i x y; reflexivity.
We can see that it discharges your lemma without any issue:
Lemma foo_eq : foo = foo'.
Proof.
unfold foo, foo'; eqFF.
Qed.
Here is a self-contained gist with all the imports and definitions.

Proof with false hypothesis in Isabelle/HOL Isar

I am trying to prove a lemma which in a certain part has a false hypothesis. In Coq I used to write "congruence" and it would get rid of the goal. However, I am not sure how to proceed in Isabelle Isar. I am trying to prove a lemma about my le function:
primrec le::"nat ⇒ nat ⇒ bool" where
"le 0 n = True" |
"le (Suc k) n = (case n of 0 ⇒ False | Suc j ⇒ le k j)"
lemma def_le: "le a b = True ⟷ (∃k. a + k = b)"
proof
assume H:"le a b = True"
show "∃k. a + k = b"
proof (induct a)
case 0
show "∃k. 0 + k = b"
proof -
have "0 + b = b" by simp
thus ?thesis by (rule exI)
qed
case Suc
fix n::nat
assume HI:"∃k. n + k = b"
show "∃k. (Suc n) + k = b"
proof (induct b)
case 0
show "∃k. (Suc n) + k = 0"
proof -
have "le (Suc n) 0 = False" by simp
oops
Note that my le function is "less or equal". At this point of the proof I find I have the hypothesis H which states that le a b = True, or in this case that le (Suc n) 0 = True which is false. How can I solve this lemma?
Another little question: I would like to write have "le (Suc n) 0 = False" by (simp only:le.simps) but this does not work. It seems I need to add some rule for reducing case expressions. What am I missing?
Thank you very much for your help.
The problem is not that it is hard to get rid of a False hypothesis in Isabelle. In fact, pretty much all of Isabelle's proof methods will instantly prove anything if there is False in the assumptions. No, the problem here is that at that point of the proof, you don't have the assumptions you need anymore, because you did not chain them into the induction. But first, allow me to make a few small remarks, and then give concrete suggestions to fix your proof.
A Few Remarks
It is somewhat unidiomatic to write le a b = True or le a b = False in Isabelle. Just write le a b or ¬le a b.
Writing the definition in a convenient form is very important to get good automation. Your definition works, of course, but I suggest the following one, which may be more natural and will give you a convenient induction rule for free:
Using the function package:
fun le :: "nat ⇒ nat ⇒ bool" where
"le 0 n = True"
| "le (Suc k) 0 = False"
| "le (Suc k) (Suc n) = le k n"
Existentials can sometimes hide important information, and they tend mess with automation, since the automation never quite knows how to instantiate them.
If you prove the following lemma, the proof is fully automatic:
lemma def_le': "le a b ⟷ a + (b - a) = b"
by (induction a arbitrary: b) (simp_all split: nat.split)
Using my function definition, it is:
lemma def_le': "le a b ⟷ (a + (b - a) = b)"
by (induction a b rule: le.induct) simp_all
Your lemma then follows from that trivially:
lemma def_le: "le a b ⟷ (∃k. a + k = b)"
using def_le' by auto
This is because the existential makes the search space explode. Giving the automation something concrete to follow helps a lot.
The actual answer
There are a number of problems. First of all, you will probably need to do induct a arbitrary: b, since the b will change during your induction (for le (Suc a) b, you will have to do a case analysis on b, and then in the case b = Suc b' you will go from le (Suc a) (Suc b') to le a b').
Second, at the very top, you have assume "le a b = True", but you do not chain this fact into the induction. If you do induction in Isabelle, you have to chain all required assumptions containing the induction variables into the induction command, or they will not be available in the induction proof. The assumption in question talks about a and b, but if you do induction over a, you will have to reason about some arbitrary variable a' that has nothing to do with a. So do e.g:
assume H:"le a b = True"
thus "∃k. a + k = b"
(and the same for the second induction over b)
Third, when you have several cases in Isar (e.g. during an induction or case analysis), you have to separate them with next if they have different assumptions. The next essentially throws away all the fixed variables and local assumptions. With the changes I mentioned before, you will need a next before the case Suc, or Isabelle will complain.
Fourth, the case command in Isar can fix variables. In your Suc case, the induction variable a is fixed; with the change to arbitrary: b, an a and a b are fixed. You should give explicit names to these variables; otherwise, Isabelle will invent them and you will have to hope that the ones it comes up with are the same as those that you use. That is not good style. So write e.g. case (Suc a b). Note that you do not have to fix variables or assume things when using case. The case command takes care of that for you and stores the local assumptions in a theorem collection with the same name as the case, e.g. Suc here. They are categorised as Suc.prems, Suc.IH, Suc.hyps. Also, the proof obligation for the current case is stored in ?case (not ?thesis!).
Conclusion
With that (and a little bit of cleanup), your proof looks like this:
lemma def_le: "le a b ⟷ (∃k. a + k = b)"
proof
assume "le a b"
thus "∃k. a + k = b"
proof (induct a arbitrary: b)
case 0
show "∃k. 0 + k = b" by simp
next
case (Suc a b)
thus ?case
proof (induct b)
case 0
thus ?case by simp
next
case (Suc b)
thus ?case by simp
qed
qed
next
It can be condensed to
lemma def_le: "le a b ⟷ (∃k. a + k = b)"
proof
assume "le a b"
thus "∃k. a + k = b"
proof (induct a arbitrary: b)
case (Suc a b)
thus ?case by (induct b) simp_all
qed simp
next
But really, I would suggest that you simply prove a concrete result like le a b ⟷ a + (b - a) = b first and then prove the existential statement using that.
Manuel Eberl did the hard part, and I just respond to your question on how to try and control simp, etc.
Before continuing, I go off topic and clarify something said on another site. The word "a tip" was used to give credit to M.E., but it should have been "3 explanations provided over 2 answers". Emails on mailing lists can't be corrected without spamming the list.
Some short answers are these:
There is no guarantee of completely controlling simp, but attributes del and only, shown below, will many times control it to the extent that you desire. To see that it's not doing more than you want, traces need to be used; an example of traces is given below.
To get complete control of proof steps, you would use "controlled" simp, along with rule, drule, and erule, and other methods. Someone else would need to give an exhaustive list.
Most anyone with the expertise to be able to answer "what's the detailed proof of what simp, auto, blast, etc does" will very rarely be willing to put in the work to answer the question. It can be plain, tedious work to investigate what simp is doing.
"Black box proofs" are always optional, as far as I can tell, if we want them to be and have the expertise to make them optional. Expertise to make them optional is generally a major limiting factor. With expertise, motivation becomes the limiting factor.
What's simp up to? It can't please everyone
If you watch, you'll see. People complain there's too much automation, or they complain there's too little automation with Isabelle.
There can never be too much automation, but that's because with Isabelle/HOL, automation is mostly optional. The possibility of no automation is what makes proving potentially interesting, but with only no automation, proving is nothing but pure tediousness, in the grand scheme.
There are attributes only and del, which can be used to mostly control simp. Speaking only from experimenting with traces, even simp will call other proof methods, similar to how auto calls simp, blast, and others.
I think you cannot prevent simp from calling linear arithmetic methods. But linear arithmetic doesn't apply much of the time.
Get set up for traces, and even the blast trace
My answer here is generalized for also trying to determine what auto is up to. One of the biggest methods that auto resorts to is blast.
You don't need the attribute_setups if you don't care about seeing when blast is used by auto, or called directly. Makarius Wenzel took the blast trace out, but then was nice enough to show the code on how to implement it.
Without the blast part, there is just the use of declare. In a proof, you can use using instead of declare. Take out what you don't want. Make sure you look at the new simp_trace_new info in the PIDE Simplifier Trace panel.
attribute_setup blast_trace = {*
Scan.lift
(Parse.$$$ "=" -- Args.$$$ "true" >> K true ||
Parse.$$$ "=" -- Args.$$$ "false" >> K false ||
Scan.succeed true) >>
(fn b => Thm.declaration_attribute (K (Config.put_generic Blast.trace b)))
*}
attribute_setup blast_stats = {*
Scan.lift
(Parse.$$$ "=" -- Args.$$$ "true" >> K true ||
Parse.$$$ "=" -- Args.$$$ "false" >> K false ||
Scan.succeed true) >>
(fn b => Thm.declaration_attribute (K (Config.put_generic Blast.stats b)))
*}
declare[[simp_trace_new mode=full]]
declare[[linarith_trace,rule_trace,blast_trace,blast_stats]]
Try and control simp, to your heart's content with only & del
I don't want to work hard by using the formula in your question. With simp, what you're looking for with only and the traces is that no rule was used that you weren't expecting.
Look at the simp trace to see what basic rewrites are done that will always be done, like basic rewrites for True and False. If you don't even want that, then you have to resort to methods like rule.
A starting point to see if you can completely shut down simp is apply(simp only:).
Here are a few examples. I would have to work harder to find an example to show when linear arithmetic is automatically being used:
lemma
"a = 0 --> a + b = (b::'a::comm_monoid_add)"
apply(simp only:) (*
ERROR: simp can't apply any magic whatsoever.
*)
oops
lemma
"a = 0 --> a + b = (b::'a::comm_monoid_add)"
apply(simp only: add_0) (*
ERROR: Still can't. Rule 'add_0' is used, but it can't be used first.
*)
oops
lemma
"a = 0 --> a + b = (b::'a::comm_monoid_add)"
apply(simp del: add_0) (*
A LITTLE MAGIC:
It applied at least one rule. See the simp trace. It tried to finish
the job automatically, but couldn't. It says "Trying to refute subgoal 1,
etc.".
Don't trust me about this, but it looks typical of blast. I was under
the impressions that simp doesn't call blast.*)
oops
lemma
"a = 0 --> a + b = (b::'a::comm_monoid_add)"
by(simp) (*
This is your question. I don't want to step through the rules that simp
uses to prove it all.
*)

Isabelle's Code generation: Abstraction lemmas for containers?

I am experimenting with the Code generator. My theory contains a datatype that encodes an invariant:
typedef small = "{x::nat. x < 10}" morphisms to_nat small
by (rule exI[where x = 0], simp)
definition "is_one x ⟷ x = small 1"
Now I want to export code for is_one. It seems that I first have to set up the data type for the code generator as follows:
code_datatype small
instantiation small :: equal
begin
definition "HOL.equal a b ⟷ to_nat a = to_nat b"
instance
apply default
unfolding equal_small_def
apply (rule to_nat_inject)
done
end
lemma [code abstype]: "small (to_nat x) = x" by (rule to_nat_inverse)
And now I am able to define a code equation for is_one that does not violate the abstraction:
definition [code del]:"small_one = small 1"
declare is_one_def[code del]
lemma [code]: "is_one x ⟷ x = small_one" unfolding is_one_def small_one_def..
lemma [code abstract]: "to_nat small_one = 1"
unfolding small_one_def by (rule small_inverse, simp)
export_code is_one in Haskell file -
Question 1: Can I avoid pulling the definition of small_one out of is_one?
Now I have a compound value of small items:
definition foo :: "small set list" where "foo = [ {small (10-i), small i} . i <- [2,3,4] ]"
In that form, I cannot export code using it (“Abstraction violation in code equation foo...”).
Question 2: How do I define an [code abstract] lemma for such a value? It seems that the code generator does not accept lemmas of the form map to_nat foo = ....
For types with invariants declared by code_abstract such as small, no code equation may contain the abstraction function small. Consequently, if you want the equality test in is_one to be expressed on type small, you have to define a constant for small 1 and prove the corresponding code equation with to_nat first. This is what you have done. However, it would be easier to use equality on the representation type nat, i.e., inline the implementation equal_class.equal; then, you do not need to instantiate equal for small either.
lemma is_one_code [code]: "is_one x ⟷ to_nat x = 1"
by(cases x)(simp add: is_one_def small_inject small_inverse)
The lifting package does most of this for you automatically:
setup_lifting type_definition_small
lift_definition is_one :: "small => bool" is "%x. x = 1" ..
export_code is_one in SML file -
Unfortunately, the code generator does not support compound return types that involve abstract types like in small set list. As described in Data Refinement in Isabelle/HOL, sect. 3.2, you have to wrap the type small set list as a type constructor small_set_list of its own and define foo with type small_set_list. The type small_set_list is then represented as nat set list with the invariant list_all (%A. ALL x:A. x < 10).
Besides Andreas' explanation and comprehensive solution, I found a work-around that avoids introducing a new type. Assume foo is not used in many places, say, only in
definition "in_foo x = (∀s ∈ set foo. x ∈ s)"
Trying to export code for in_foo gives the "Abstraction violation" error. But using program refinement, I can avoid the problematic code:
I use code_thms to find out how foo is used and find that it is used as list_all (op ∈ ?x) foo. Now I rewrite this particular use of foo without mentioning small:
lemma [code_unfold]:
"list_all (op ∈ x) foo ⟷
list_all (op ∈ (to_nat x)) [ {(10-i), i} . i <- [2, 3, 4] ]"
unfolding foo_def by (auto simp add: small_inverse small_inject)
(Had I defined foo using lift_definition, the proof after apply transfer would be even simpler.)
After this transformation, which essentially captures that all members of foo fulfill the invariant of small, the code export works as intended:
export_code in_foo in Haskell file -