My aim is to prove properties of lists containing generated patterns.
In the first example the pattern is simply a sequence of 0s and lemma pattern_0_len proves that the length of the generated list indeed equals to the length parameter of the generator function.
theory pattern_0
imports Main
begin
fun pattern_0 :: "nat ⇒ nat list" where
"pattern_0 0 = []" |
"pattern_0 len = (pattern_0 (len - 1)) # [0]"
lemma pattern_0_len [simp]: "length (pattern_0 lng) = lng"
apply(induction lng)
apply(simp)
apply(auto)
done
end
In the second example the generator produces a sequence of 0, 1 items.
theory pattern_0_1
imports Main
begin
fun pattern_0_1 :: "nat ⇒ nat ⇒ nat list" where
"pattern_0_1 0 item = []" |
"pattern_0_1 len item = (pattern_0_1 (len - 1) (if item = 0 then 1 else 0)) # [item]"
lemma pattern_0_1_len [simp]: "length (pattern_0_1 lng item) = lng"
apply(induction lng)
apply(simp)
apply(auto)
done
end
Unfortunately, pattern_0_1_len does not prove (after simp the goal is exactly the induction step) and I'd like to understand the reason why not. Is it the presence of the item parameter that 'confuses' Isabelle? What can be done in this situation, preferably without declaring anything about how the pattern is generated?
The additional parameter is indeed the problem. For example, consider this subgoal:
1. ⋀lng. length (pattern_0_1 lng 0) = lng ⟹ item = 0 ⟹ length (pattern_0_1 lng (Suc 0)) = lng
You see that the induction hypothesis is only applicable for zero, but you need it for one.
The fix is simple:
apply(induction lng arbitrary: item)
This instructs the induction method to first generalize the variable item. Then, the induction hypothesis becomes more broadly applicable.
Related
I have datatype stack_op which consists of several (~20) cases. I'm trying write function which skips some of that cases in list:
function (sequential) skip_expr :: "stack_op list ⇒ stack_op list" where
"skip_expr [] = []"
| "skip_expr ((stack_op.Unary _)#other) = (skip_expr other)"
| "skip_expr ((stack_op.Binary _)#other) = skip_expr (skip_expr other)"
| "skip_expr ((stack_op.Value _)#other) = other"
| "skip_expr other = other"
by pat_completeness auto termination by lexicographic_order
which seems to always terminate. But trying by lexicographic order generates such unresolved cases:
Calls:
c) stack_op.Binary uv_ # other ~> skip_expr other
Measures:
1) size_list size
2) length
Result matrix:
1 2
c: ? ?
(size_change also desn't work)
I've read https://isabelle.in.tum.de/dist/Isabelle2021/doc/functions.pdf, but it couldn't help. (Maybe there are more complex examples of tremination use?)
I tried to rewrite function adding another param:
function (sequential) skip_expr :: "stack_op list ⇒ nat ⇒ stack_op list" where
"skip_expr l 0 = l"
| "skip_expr [] _ = []"
| "skip_expr ((stack_op.Unary _)#other) depth = (skip_expr other (depth - 1))"
| "skip_expr ((stack_op.Binary _)#other) depth =
(let buff1 = (skip_expr other (depth - 1))
in (skip_expr buff1 (length buff1)))"
| "skip_expr ((stack_op.Value _)#other) _ = other"
| "skip_expr other _ = other"
by pat_completeness auto
termination by (relation "measure (λ(_,dep). dep)") auto
which generates unresolved subgoal:
1. ⋀other v. skip_expr_dom (other, v) ⟹ length (skip_expr other v) < Suc v
which I also don't how to proof.
Could anyone how such cases solved (As I can understand there is some problem with two-level recursive call on rigth side of stack_op.Binary case)? Or maybe there is another way to make such skip?
Thanks in advance
The lexicographic_order method simply tries to solve the arising goals with the simplifier, so if the simplifier gets stuck you end up with unresolved termination subgoals.
In this case, as you identified correctly, the problem is that you have a nested recursive call skip_expr (skip_expr other). This is always problematic because at this stage, the simplifier knows nothing about what skip_expr does to the input list. For all we know, it might just return the list unmodified, or even a longer list, and then it surely would not terminate.
Confronting the issue head on
The solution is to show something about length (skip_expr …) and make that information available to the simplifier. Because we have not yet shown termination of the function, we have to use the skip_expr.psimps rules and the partial induction rule skip_expr.pinduct, i.e. every statement we make about skip_expr xs always has as a precondition that skip_expr actually terminates on the input xs. For this, there is the predicate skip_expr_dom.
Putting it all together, it looks like this:
lemma length_skip_expr [termination_simp]:
"skip_expr_dom xs ⟹ length (skip_expr xs) ≤ length xs"
by (induction xs rule: skip_expr.pinduct) (auto simp: skip_expr.psimps)
termination skip_expr by lexicographic_order
Circumventing the issue
Sometimes it can also be easier to circumvent the issue entirely. In your case, you could e.g. define a more general function skip_exprs that skips not just one instruction but n instructions. This you can define without nested induction:
fun skip_exprs :: "nat ⇒ stack_op list ⇒ stack_op list" where
"skip_exprs 0 xs = xs"
| "skip_exprs (Suc n) [] = []"
| "skip_exprs (Suc n) (Unary _ # other) = skip_exprs (Suc n) other"
| "skip_exprs (Suc n) (Binary _ # other) = skip_exprs (Suc (Suc n)) other"
| "skip_exprs (Suc n) (Value _ # other) = skip_exprs n other"
| "skip_exprs (Suc n) xs = xs"
Equivalence to your skip_expr is then straightforward to prove:
lemma skip_exprs_conv_skip_expr: "skip_exprs n xs = (skip_expr ^^ n) xs"
proof -
have [simp]: "(skip_expr ^^ n) [] = []" for n
by (induction n) auto
have [simp]: "(skip_expr ^^ n) (Other # xs) = Other # xs" for xs n
by (induction n) auto
show ?thesis
by (induction n xs rule: skip_exprs.induct)
(auto simp del: funpow.simps simp: funpow_Suc_right)
qed
lemma skip_expr_Suc_0 [simp]: "skip_exprs (Suc 0) xs = skip_expr xs"
by (simp add: skip_exprs_conv_skip_expr)
In your case, I don't think it actually makes sense to do this because figuring out the termination is fairly easy, but it may be good to keep in mind.
This pattern generator produces a list with a given number at a given position, all other values are zero.
fun pattern_one_value :: "nat ⇒ nat ⇒ nat ⇒ nat ⇒ nat list" where
"pattern_one_value _ _ _ 0 = []" |
"pattern_one_value pos pos1 val lng =
(if pos = pos1 then val else 0) # (pattern_one_value pos (pos1 + 1) val (lng - 1))"
The following lemma is aimed to prove that generated lists contain the right value at the given position.
lemma pattern_one_value_check [simp]: "∀pos val. pos < lng ⟹ pattern_one_value pos 0 val lng ! pos = val"
proof(induct lng)
case 0 then show ?case by simp
next
case (Suc lng) then show ?case by auto
qed
It seems to be a correct proof; however, changing val in the cons expression of the generator function into an arbitrary number like (if pos = pos1 then 7 else 0) # ..., the proof still holds because both the base and the induction hypothesis are false.
Where am I wrong? Thanks for any help.
It seems to be a correct proof; however, changing val in the cons
expression of the generator function into an arbitrary number like (if pos = pos1 then 7 else 0) # ..., the proof still holds because both
the base and the induction hypothesis are false. Where am I wrong?
I believe that the problem is related to an attempt to treat HOL's universal quantifier ∀ as equivalent to Pure's universal quantifier ⋀. Effectively, it is possible to prove anything from the premise of the theorem pattern_one_value_check, as stated in your question. Indeed:
lemma pattern_one_value_check'[simp]:
"(∀pos val::nat. pos < (lng::nat)) = False"
by auto
lemma pattern_one_value_check''[simp]:
"(∀pos val::nat. pos < (lng::nat)) ⟹ P"
by auto
I believe that you meant to use Pure's universal quantification in the statement of the theorem, e.g.
lemma pattern_one_value_check [simp]:
"⋀pos val. pos < lng ⟹ pattern_one_value pos 0 val lng ! pos = val"
proof(induct lng)
case 0 then show ?case by simp
next
case (Suc lng) then show ?case sorry
qed
In fact, even this is not necessary. The following theorem, once proven, will appear in the context as identical to the one stated above:
lemma pattern_one_value_check' [simp]:
"pos < lng ⟹ pattern_one_value pos 0 val lng ! pos = val"
proof(induct lng)
case 0 then show ?case by simp
next
case (Suc lng) then show ?case sorry
qed
If you seek a more detailed explanation, see Section 2.1 in Isar-ref and the document "Programming and Proving in Isabelle/HOL", both are part of the official documentation.
As a side note, I have to mention that, perhaps, there is an easier way to define pattern_one_value. In this case, the proof of pattern_one_value_check also seems to be easier:
definition pattern_one_value :: "nat ⇒ nat ⇒ nat ⇒ nat list"
where "pattern_one_value val pos len = list_update (replicate len 0) pos val"
lemma pattern_one_value_check:
assumes "pos < len"
shows "pattern_one_value val pos len ! pos = val"
using assms unfolding pattern_one_value_def
apply(induct len)
subgoal by auto
subgoal by (metis length_replicate nth_list_update)
done
Given a function that generates a list of identical items I wish to prove that the generated lists consist the given natural number at all positions independent of list length.
fun pattern_n :: "nat ⇒ nat ⇒ nat list" where
"pattern_n _ 0 = []" |
"pattern_n n lng = n # (pattern_n n (lng - 1))"
lemma pattern_n_1: "lng > 0 ∧ pos ≥ 0 ∧ pos < lng ∧ n ≥ 0 ⟹ (pattern_n n lng ! pos) = n"
It seems obvious that the proof should be based on induction on the length of the generated list but pos also seems to be an induction variable candidate. I'd appreciate any help on how to proceed with this proof.
The function pattern_n is equivalent to the function replicate from the standard library (theory List). The standard library also contains the theorem nth_replicate for the function replicate that is nearly identical to the theorem that you are trying to prove:
fun pattern_n :: "nat ⇒ nat ⇒ nat list" where
"pattern_n _ 0 = []" |
"pattern_n n lng = n # (pattern_n n (lng - 1))"
lemma "pattern_n n k = replicate k n"
by (induction k) auto
thm nth_replicate
UPDATE
Alternatively, you can use induction to prove the result. Usually it is more convenient to use the definition in the form that is provided by the function pattern_n' below, because the theorems that are generated automatically when you define the function are more consistent with this form.
fun pattern_n :: "nat ⇒ nat ⇒ nat list" where
"pattern_n _ 0 = []" |
"pattern_n n lng = n # (pattern_n n (lng - 1))"
fun pattern_n' :: "nat ⇒ nat ⇒ nat list" where
"pattern_n' n 0 = []" |
"pattern_n' n (Suc lng) = n # (pattern_n' n lng)"
lemma "pattern_n n lng = pattern_n' n lng"
by (induct lng) auto
lemma pattern_n_1_via_replicate:
"pos < lng ⟹ (pattern_n val lng) ! pos = val"
proof(induct lng arbitrary: pos)
case 0 then show ?case by simp
next
case (Suc lng) then show ?case by (fastforce simp: less_Suc_eq_0_disj)
qed
Isabelle version: Isabelle2020
I'm a mathematician just starting to get used to Isabelle, and something that should be incredibly simple turned out to be frustrating. How do I define a function between two constants? Say, the function f: {1,2,3} \to {1,2,4} mapping 1 to 1, 2 to 4 and 3 to 2?
I suppose I managed to define the sets as constants t1 and t2 without incident, but (I guess since they're not datatypes) I can't try something like
definition f ::"t1 => t2" where
"f 1 = 1" |
"f 2 = 4" |
"f 3 = 2"
I believe there must be a fundamental misconception behind this difficulty, so I appreciate any guidance.
There's a number of aspects to your question.
First, to get something working quickly, use the fun keyword instead of definition, like so:
fun test :: "nat ⇒ nat" where
"test (Suc 0) = 1" |
"test (Suc (Suc 0)) = 4" |
"test (Suc (Suc (Suc 0))) = 2" |
"test _ = undefined"
You cannot pattern match on any arguments directly in the head of the definition using the definition keyword, whereas you can with fun. Note also that I have replaced the overloaded numeric literals (1, 2, 3, etc.) with the constructors for the nat datatype (0 and Suc) in the pattern match.
An alternative would be to stick with definition, but push the case analysis of the function's argument inside the body of the definition using a case statement, like so:
definition test2 :: "nat ⇒ nat" where
"test2 x ≡
case x of
(Suc 0) ⇒ 1
| (Suc (Suc 0)) ⇒ 4
| (Suc (Suc (Suc 0))) ⇒ 2
| _ ⇒ undefined"
Note that definitions like test2 are not unfolded by the simplifier by default, and you will need to manually add the theorem test2_def to the simplifier's simpset if you want to expand occurrences of test2 in a proof.
You can also define new types (you cannot use sets as types, directly, as you are trying to do) corresponding to your two three-element sets with typedef, but personally I would stick with nat.
EDIT: to use typedef, do something like:
typedef t1 = "{x::nat. x = 1 ∨ x = 2 ∨ x = 3}"
by auto
definition test :: "t1 ⇒ t1" where
"test x ≡
case (Rep_t1 x) of
| Suc 0 ⇒ Abs_t1 1
| Suc (Suc 0) ⇒ Abs_t1 4
| Suc (Suc (Suc 0)) ⇒ Abs_t1 2"
Though, I don't really ever use typedef myself, and so this may not be the best way of using this and others may possibly suggest some other way. What typedef does is carve out a new type from an existing one, by identifying a non-empty set of inhabitants for the new type. The proof obligation, here closed by auto, is merely to demonstrate that the defining set for the new type is indeed non-empty, and in this case I am carving out a three-element set of naturals into a new type, called t1, so the proof is fairly trivial. Two new constants are created, Abs_t1 and Rep_t1 which allow you to move back-and-forth between the naturals and the new type. If you put a print_theorems after the typedef command you will see several new theorems about t1 that Isabelle has automatically generated for you.
Up until several days ago, I always defined a type, and then proved theorems directly about the type. Now I'm trying to use type classes.
Problem
The problem is that I can't instantiate cNAT for my type myD below, and it appears it's because simp has no effect on the abstract function cNAT, which I've made concrete with my primrec function cNAT_myD. I can only guess what's happening because of the automation that happens after instance proof.
Questions
Q1: Below, at the statement instantiation myD :: (type) cNAT, can you tell me how to finish the proof, and why I can prove the following theorem, but not the type class proof, which requires injective?
theorem dNAT_1_to_1: "(dNAT n = dNAT m) ==> n = m"
assumes injective: "(cNAT n = cNAT m) ==> n = m"
Q2: This is not as important, but at the bottom is this statement:
instantiation myD :: (type) cNAT2
It involves another way I was trying to instantiate cNAT. Can you tell me why I get Failed to refine any pending goal at shows? I put some comments in the source to explain some of what I did to set it up. I used this slightly modified formula for the requirement injective:
assumes injective: "!!n m. (cNAT2 n = cNAT2 m) --> n = m"
Specifics
My contrived datatype is this, which may be useful to me someday: (Update: Well, for another example maybe. A good mental exercise is for me to try and figure out how I can actually get something inside a 'a myD list, other than []. With BNF, something like datatype_new 'a myD = myS "'a myD fset" gives me the warning that there's an unused type variable on the right-hand side)
datatype 'a myD = myL "'a myD list"
The type class is this, which requires an injective function from nat to 'a:
class cNAT =
fixes cNAT :: "nat => 'a"
assumes injective: "(cNAT n = cNAT m) ==> n = m"
dNAT: this non-type class version of cNAT works
fun get_myL :: "'a myD => 'a myD list" where
"get_myL (myL L) = L"
primrec dNAT :: "nat => 'a myD" where
"dNAT 0 = myL []"
|"dNAT (Suc n) = myL (myL [] # get_myL(dNAT n))"
fun myD2nat :: "'a myD => nat" where
"myD2nat (myL []) = 0"
|"myD2nat (myL (x # xs)) = Suc(myD2nat (myL xs))"
theorem left_inverse_1 [simp]:
"myD2nat(dNAT n) = n"
apply(induct n, auto)
by(metis get_myL.cases get_myL.simps)
theorem dNAT_1_to_1:
"(dNAT n = dNAT m) ==> n = m"
apply(induct n)
apply(simp) (*
The simp method expanded dNAT.*)
apply(metis left_inverse_1 myD2nat.simps(1))
by (metis left_inverse_1)
cNAT: type class version that I can't instantiate
instantiation myD :: (type) cNAT
begin
primrec cNAT_myD :: "nat => 'a myD" where
"cNAT_myD 0 = myL []"
|"cNAT_myD (Suc n) = myL (myL [] # get_myL(cNAT_myD n))"
instance
proof
fix n m :: nat
show "cNAT n = cNAT m ==> n = m"
apply(induct n)
apply(simp) (*
The simp method won't expand cNAT to cNAT_myD's definition.*)
by(metis injective)+ (*
Metis proved it without unfolding cNAT_myD. It's useless. Goals always remain,
and the type variables in the output panel are all weird.*)
oops
end
cNAT2: Failed to refine any pending goal at show
(*I define a variation of `injective` in which the `assumes` definition, the
goal, and the `show` statement are exactly the same, and that strange `fails
to refine any pending goal shows up.*)
class cNAT2 =
fixes cNAT2 :: "nat => 'a"
assumes injective: "!!n m. (cNAT2 n = cNAT2 m) --> n = m"
instantiation myD :: (type) cNAT2
begin
primrec cNAT2_myD :: "nat => 'a myD" where
"cNAT2_myD 0 = myL []"
|"cNAT2_myD (Suc n) = myL (myL [] # get_myL(cNAT2_myD n))"
instance
proof (*
goal: !!n m. cNAT2 n = cNAT2 m --> n = m.*)
show
"!!n m. cNAT2 n = cNAT2 m --> n = m"
(*Failed to refine any pending goal
Local statement fails to refine any pending goal
Failed attempt to solve goal by exported rule:
cNAT2 (n::nat) = cNAT2 (m::nat) --> n = m *)
Your function cNAT is polymorphic in its result type, but the type variable does not appear among the parameters. This often causes type inference to compute a type which is more general than you want. In your case for cNAT, Isabelle infers for the two occurrences of cNAT in the show statement the type nat => 'b for some 'b of sort cNAT, but their type in the goal is nat => 'a myD. You can see this in jEdit by Ctrl-hovering over the cNAT occurrences to inspect the types. In ProofGeneral, you can enable printing of types with using [[show_consts]].
Therefore, you have to explicitly constrain types in the show statement as follows:
fix n m
assume "(cNAT n :: 'a myD) = cNAT m"
then show "n = m"
Note that it is usually not a good idea to use Isabelle's meta-connectives !! and ==> inside a show statement, you better rephrase them using fix/assume/show.