Stream: Archive Mirror: Isabelle Users Mailing List

Topic: [isabelle] a proof on primitive recursion


view this post on Zulip Email Gateway (Aug 22 2022 at 13:49):

From: Gergely Buday <buday.gergely@uni-eszterhazy.hu>
Hi,

I have proved a simple associativity rule and tried to prove commutativity of an addition function:

theory Add

imports Main

begin

datatype natural = Zero | Succ natural

primrec add :: "natural ⇒ natural ⇒ natural"

where

"add Zero m = m"

| "add (Succ n) m = Succ (add n m)"

lemma add_assoc: "⋀ k m n . add k (add m n) = add (add k m) n "

proof -

fix k m n

show "add k (add m n) = add (add k m) n"

proof (induct k arbitrary: m n)

fix m n

show "add Zero (add m n) = add (add Zero m) n"

proof (induct m arbitrary: n)

fix n

show "add Zero (add Zero n) = add (add Zero Zero) n" by simp

next

fix m n

assume "⋀n. add Zero (add m n) = add (add Zero m) n"

from this show "add Zero (add (Succ m) n) = add (add Zero (Succ m)) n" by simp

qed

next

fix k m n

assume "⋀m n. add k (add m n) = add (add k m) n"

from this show "add (Succ k) (add m n) = add (add (Succ k) m) n" by simp

qed

qed

text{*

Questions:

*}

lemma add_commute: "⋀ m n. add m n = add n m"

proof -

fix m n

show "add m n = add n m"

proof (induct m arbitrary: n)

fix n

show "add Zero n = add n Zero"

proof (induct n)

show "add Zero Zero = add Zero Zero" by simp

next

fix n

assume "add Zero n = add n Zero"

from this show "add Zero (Succ n) = add (Succ n) Zero" by simp

qed

next

fix m n

assume indhyp: "⋀n. add m n = add n m"

from this show "add (Succ m) n = add n (Succ m)"


I am stuck here. There is no way of rewriting Succ m in the second parameter of Succ. What is the way out of the bottle?

Should I try another definition?

Should I use the proven theory of builtin natural numbers, prove an adequacy theorem of + and this add function and use the commutativity of + ?, or I can make it without this recourse?

view this post on Zulip Email Gateway (Aug 22 2022 at 13:49):

From: Alfio Martini <alfio.martini@acm.org>
Hi Gergely,

You need an auxiliary lemma. Using a slightly different notation. Here is a
very old version.
I would not do it in this way nowadays :-)

Best!

lemma lem01:"add zero n = n"
apply (induction n)
apply (simp)
apply (simp)
done

This is the auxiliary lemma!

lemma lem02: "∀x. suc (add x n) = add (suc x) n"
apply (induction n)
apply (simp)
apply (simp)
done

theorem th02a:"∀x. add x n = add n x"
apply (induction n)
apply (simp_all add:lem01 lem02)
done

Best!


Last updated: Apr 25 2024 at 20:15 UTC