Stream: Archive Mirror: Isabelle Users Mailing List

Topic: [isabelle] case rule and OF


view this post on Zulip Email Gateway (Aug 22 2022 at 14:08):

From: Esseger <esseger@free.fr>
Dear all,

I am trying to set a convenient case rule, to distinguish if, given a
parameter p ≥ 1 in ennreal, it is equal to 1, or between 1 and infinity,
or infinite. When it is strictly between 1 and infinity,
I would also like p to be written as p = ennreal p2, for some real
number p2.

I would like to use it as:

lemma
assumes "p ≥ (1::ennreal)"
show foo
proof (rule my_case_rule[OF p ≥ 1])
case one
then show ?thesis sorry
next
case (gr p2)
then show ?thesis sorry
next
case PInf
then show ?thesis sorry
qed

I tried to define my rule as:

lemma my_case_rule:
assumes "p ≥ (1::ennreal)"
obtains (gr) p2 where "p = ennreal p2" "p2 > 1"
| (one) "p = 1"
| (PInf) "p = ∞"
using assms by (metis (full_types) antisym_conv ennreal_cases
ennreal_le_1 infinity_ennreal_def not_le)

It doesn't work, as the names of the cases are not respected, replacing
them with the names 1, 2, 3 (with which I can write the proofs, for
sure, but this is much harder to read).

I tried several modifications of the rule to keep the names, with
modifiers such as case_names or consumes, to no avail. I could not
locate in the reference manual any hint on the canonical way to write
this kind of thing. Any help on the best practice in this situation?

Best,
Esseger

view this post on Zulip Email Gateway (Aug 22 2022 at 14:08):

From: Esseger <esseger@free.fr>
Sorry, of course I meant:

proof (cases rule: my_case_rule[OF p ≥ 1])

My problem is that the cases names are not respected with this construction.

view this post on Zulip Email Gateway (Aug 22 2022 at 14:08):

From: Andreas Lochbihler <andreas.lochbihler@inf.ethz.ch>
Dear Esseger,

When you instantiate some premises of a rule with OF, all the case name information gets
erased (because it is not trivial to determine how they would have to be shifted).
However, you normally do not instantiate premises manually in case distinctions, but
appropriately declare the first n premises to be unified with the facts chained in using
the "consumes n" attribute. "lemma assumes ... obtains ..." implicitly sets "consumes"
appropriately. Hence, you merely have to chain in the necessary assumptions, as I had
shown in my example.

Andreas

view this post on Zulip Email Gateway (Aug 22 2022 at 14:08):

From: Andreas Lochbihler <andreas.lochbihler@inf.ethz.ch>
Dear Esseger,

Case names as specified for rules are only honored by the proof methods "cases", "induct"
and "coinduct" (and their derivatives "induction" and "coinduction"), but not plain
"rule". Your rule "my_case_rule" should work if you use "cases" (which is the appropriate
method here, because you are doing a case distinction). So, the following should do the job:

lemma
assumes "p ≥ (1::ennreal)"
show foo
using p >= 1
proof (cases rule: my_case_rule)

Hope this helps,
Andreas

view this post on Zulip Email Gateway (Aug 22 2022 at 14:08):

From: Esseger <esseger@free.fr>
This is perfectly clear, thanks a lot for the explanations.


Last updated: Apr 25 2024 at 20:15 UTC