Theory PRecFun2
section ‹Primitive recursive functions of one variable›
theory PRecFun2
imports PRecFun
begin
subsection ‹Alternative definition of primitive recursive functions of one variable›
definition
UnaryRecOp :: "(nat ⇒ nat) ⇒ (nat ⇒ nat) ⇒ (nat ⇒ nat)" where
"UnaryRecOp = (λ g h. pr_conv_2_to_1 (PrimRecOp g (pr_conv_1_to_3 h)))"
lemma unary_rec_into_pr: "⟦ g ∈ PrimRec1; h ∈ PrimRec1 ⟧ ⟹ UnaryRecOp g h ∈ PrimRec1" by (simp add: UnaryRecOp_def pr_conv_1_to_3_lm pr_conv_2_to_1_lm pr_rec)
definition
c_f_pair :: "(nat ⇒ nat) ⇒ (nat ⇒ nat) ⇒ (nat ⇒ nat)" where
"c_f_pair = (λ f g x. c_pair (f x) (g x))"
lemma c_f_pair_to_pr: "⟦ f ∈ PrimRec1; g ∈ PrimRec1 ⟧ ⟹ c_f_pair f g ∈ PrimRec1"
unfolding c_f_pair_def by prec
inductive_set PrimRec1' :: "(nat ⇒ nat) set"
where
zero: "(λ x. 0) ∈ PrimRec1'"
| suc: "Suc ∈ PrimRec1'"
| fst: "c_fst ∈ PrimRec1'"
| snd: "c_snd ∈ PrimRec1'"
| comp: "⟦ f ∈ PrimRec1'; g ∈ PrimRec1' ⟧ ⟹ (λ x. f (g x)) ∈ PrimRec1'"
| pair: "⟦ f ∈ PrimRec1'; g ∈ PrimRec1' ⟧ ⟹ c_f_pair f g ∈ PrimRec1'"
| un_rec: "⟦ f ∈ PrimRec1'; g ∈ PrimRec1' ⟧ ⟹ UnaryRecOp f g ∈ PrimRec1'"
lemma primrec'_into_primrec: "f ∈ PrimRec1' ⟹ f ∈ PrimRec1"
proof (induct f rule: PrimRec1'.induct)
case zero show ?case by (rule pr_zero)
next
case suc show ?case by (rule pr_suc)
next
case fst show ?case by (rule c_fst_is_pr)
next
case snd show ?case by (rule c_snd_is_pr)
next
case comp from comp show ?case by (simp add: pr_comp1_1)
next
case pair from pair show ?case by (simp add: c_f_pair_to_pr)
next
case un_rec from un_rec show ?case by (simp add: unary_rec_into_pr)
qed
lemma pr_id1_1': "(λ x. x) ∈ PrimRec1'"
proof -
have "c_f_pair c_fst c_snd ∈ PrimRec1'" by (simp add: PrimRec1'.fst PrimRec1'.snd PrimRec1'.pair)
moreover have "c_f_pair c_fst c_snd = (λ x. x)" by (simp add: c_f_pair_def)
ultimately show ?thesis by simp
qed
lemma pr_id2_1': "pr_conv_2_to_1 (λ x y. x) ∈ PrimRec1'" by (simp add: pr_conv_2_to_1_def PrimRec1'.fst)
lemma pr_id2_2': "pr_conv_2_to_1 (λ x y. y) ∈ PrimRec1'" by (simp add: pr_conv_2_to_1_def PrimRec1'.snd)
lemma pr_id3_1': "pr_conv_3_to_1 (λ x y z. x) ∈ PrimRec1'"
proof -
have "pr_conv_3_to_1 (λ x y z. x) = (λx. c_fst (c_fst x))" by (simp add: pr_conv_3_to_1_def)
moreover from PrimRec1'.fst PrimRec1'.fst have "(λx. c_fst (c_fst x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
ultimately show ?thesis by simp
qed
lemma pr_id3_2': "pr_conv_3_to_1 (λ x y z. y) ∈ PrimRec1'"
proof -
have "pr_conv_3_to_1 (λ x y z. y) = (λx. c_snd (c_fst x))" by (simp add: pr_conv_3_to_1_def)
moreover from PrimRec1'.snd PrimRec1'.fst have "(λx. c_snd (c_fst x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
ultimately show ?thesis by simp
qed
lemma pr_id3_3': "pr_conv_3_to_1 (λ x y z. z) ∈ PrimRec1'"
proof -
have "pr_conv_3_to_1 (λ x y z. z) = (λx. c_snd x)" by (simp add: pr_conv_3_to_1_def)
thus ?thesis by (simp add: PrimRec1'.snd)
qed
lemma pr_comp2_1': "⟦ pr_conv_2_to_1 f ∈ PrimRec1'; g ∈ PrimRec1'; h ∈ PrimRec1' ⟧ ⟹ (λ x. f (g x) (h x)) ∈ PrimRec1'"
proof -
assume A1: "pr_conv_2_to_1 f ∈ PrimRec1'"
assume A2: "g ∈ PrimRec1'"
assume A3: "h ∈ PrimRec1'"
let ?f1 = "pr_conv_2_to_1 f"
have S1: "(%x. ?f1 ((c_f_pair g h) x)) = (λ x. f (g x) (h x))" by (simp add: c_f_pair_def pr_conv_2_to_1_def)
from A2 A3 have S2: "c_f_pair g h ∈ PrimRec1'" by (rule PrimRec1'.pair)
from A1 S2 have S3: "(%x. ?f1 ((c_f_pair g h) x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
with S1 show ?thesis by simp
qed
lemma pr_comp3_1': "⟦ pr_conv_3_to_1 f ∈ PrimRec1'; g ∈ PrimRec1'; h ∈ PrimRec1'; k ∈ PrimRec1' ⟧ ⟹ (λ x. f (g x) (h x) (k x)) ∈ PrimRec1'"
proof -
assume A1: "pr_conv_3_to_1 f ∈ PrimRec1'"
assume A2: "g ∈ PrimRec1'"
assume A3: "h ∈ PrimRec1'"
assume A4: "k ∈ PrimRec1'"
from A2 A3 have "c_f_pair g h ∈ PrimRec1'" by (rule PrimRec1'.pair)
from this A4 have "c_f_pair (c_f_pair g h) k ∈ PrimRec1'" by (rule PrimRec1'.pair)
from A1 this have "(%x. (pr_conv_3_to_1 f) ((c_f_pair (c_f_pair g h) k) x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
then show ?thesis by (simp add: c_f_pair_def pr_conv_3_to_1_def)
qed
lemma pr_comp1_2': "⟦ f ∈ PrimRec1'; pr_conv_2_to_1 g ∈ PrimRec1' ⟧ ⟹ pr_conv_2_to_1 (λ x y. f (g x y)) ∈ PrimRec1'"
proof -
assume "f ∈ PrimRec1'"
and "pr_conv_2_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
then show ?thesis by (simp add: pr_conv_2_to_1_def)
qed
lemma pr_comp1_3': "⟦ f ∈ PrimRec1'; pr_conv_3_to_1 g ∈ PrimRec1' ⟧ ⟹ pr_conv_3_to_1 (λ x y z. f (g x y z)) ∈ PrimRec1'"
proof -
assume "f ∈ PrimRec1'"
and "pr_conv_3_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x)) ∈ PrimRec1'" by (rule PrimRec1'.comp)
then show ?thesis by (simp add: pr_conv_3_to_1_def)
qed
lemma pr_comp2_2': "⟦ pr_conv_2_to_1 f ∈ PrimRec1'; pr_conv_2_to_1 g ∈ PrimRec1'; pr_conv_2_to_1 h ∈ PrimRec1' ⟧ ⟹ pr_conv_2_to_1 (λ x y. f (g x y) (h x y)) ∈ PrimRec1'"
proof -
assume "pr_conv_2_to_1 f ∈ PrimRec1'"
and "pr_conv_2_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
and "pr_conv_2_to_1 h ∈ PrimRec1'" (is "?h1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x) (?h1 x)) ∈ PrimRec1'" by (rule pr_comp2_1')
then show ?thesis by (simp add: pr_conv_2_to_1_def)
qed
lemma pr_comp2_3': "⟦ pr_conv_2_to_1 f ∈ PrimRec1'; pr_conv_3_to_1 g ∈ PrimRec1'; pr_conv_3_to_1 h ∈ PrimRec1' ⟧ ⟹ pr_conv_3_to_1 (λ x y z. f (g x y z) (h x y z)) ∈ PrimRec1'"
proof -
assume "pr_conv_2_to_1 f ∈ PrimRec1'"
and "pr_conv_3_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
and "pr_conv_3_to_1 h ∈ PrimRec1'" (is "?h1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x) (?h1 x)) ∈ PrimRec1'" by (rule pr_comp2_1')
then show ?thesis by (simp add: pr_conv_3_to_1_def)
qed
lemma pr_comp3_2': "⟦ pr_conv_3_to_1 f ∈ PrimRec1'; pr_conv_2_to_1 g ∈ PrimRec1'; pr_conv_2_to_1 h ∈ PrimRec1'; pr_conv_2_to_1 k ∈ PrimRec1' ⟧ ⟹ pr_conv_2_to_1 (λ x y. f (g x y) (h x y) (k x y)) ∈ PrimRec1'"
proof -
assume "pr_conv_3_to_1 f ∈ PrimRec1'"
and "pr_conv_2_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
and "pr_conv_2_to_1 h ∈ PrimRec1'" (is "?h1 ∈ PrimRec1'")
and "pr_conv_2_to_1 k ∈ PrimRec1'" (is "?k1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x) (?h1 x) (?k1 x)) ∈ PrimRec1'" by (rule pr_comp3_1')
then show ?thesis by (simp add: pr_conv_2_to_1_def)
qed
lemma pr_comp3_3': "⟦ pr_conv_3_to_1 f ∈ PrimRec1'; pr_conv_3_to_1 g ∈ PrimRec1'; pr_conv_3_to_1 h ∈ PrimRec1'; pr_conv_3_to_1 k ∈ PrimRec1' ⟧ ⟹ pr_conv_3_to_1 (λ x y z. f (g x y z) (h x y z) (k x y z)) ∈ PrimRec1'"
proof -
assume "pr_conv_3_to_1 f ∈ PrimRec1'"
and "pr_conv_3_to_1 g ∈ PrimRec1'" (is "?g1 ∈ PrimRec1'")
and "pr_conv_3_to_1 h ∈ PrimRec1'" (is "?h1 ∈ PrimRec1'")
and "pr_conv_3_to_1 k ∈ PrimRec1'" (is "?k1 ∈ PrimRec1'")
then have "(λ x. f (?g1 x) (?h1 x) (?k1 x)) ∈ PrimRec1'" by (rule pr_comp3_1')
then show ?thesis by (simp add: pr_conv_3_to_1_def)
qed
lemma lm': "(f1 ∈ PrimRec1 ⟶ f1 ∈ PrimRec1') ∧ (g1 ∈ PrimRec2 ⟶ pr_conv_2_to_1 g1 ∈ PrimRec1') ∧ (h1 ∈ PrimRec3 ⟶ pr_conv_3_to_1 h1 ∈ PrimRec1')"
proof (induct rule: PrimRec1_PrimRec2_PrimRec3.induct)
case zero show ?case by (rule PrimRec1'.zero)
next case suc show ?case by (rule PrimRec1'.suc)
next case id1_1 show ?case by (rule pr_id1_1')
next case id2_1 show ?case by (rule pr_id2_1')
next case id2_2 show ?case by (rule pr_id2_2')
next case id3_1 show ?case by (rule pr_id3_1')
next case id3_2 show ?case by (rule pr_id3_2')
next case id3_3 show ?case by (rule pr_id3_3')
next case comp1_1 from comp1_1 show ?case by (simp add: PrimRec1'.comp)
next case comp1_2 from comp1_2 show ?case by (simp add: pr_comp1_2')
next case comp1_3 from comp1_3 show ?case by (simp add: pr_comp1_3')
next case comp2_1 from comp2_1 show ?case by (simp add: pr_comp2_1')
next case comp2_2 from comp2_2 show ?case by (simp add: pr_comp2_2')
next case comp2_3 from comp2_3 show ?case by (simp add: pr_comp2_3')
next case comp3_1 from comp3_1 show ?case by (simp add: pr_comp3_1')
next case comp3_2 from comp3_2 show ?case by (simp add: pr_comp3_2')
next case comp3_3 from comp3_3 show ?case by (simp add: pr_comp3_3')
next case prim_rec
fix g h assume A1: "g ∈ PrimRec1'" and "pr_conv_3_to_1 h ∈ PrimRec1'"
then have "UnaryRecOp g (pr_conv_3_to_1 h) ∈ PrimRec1'" by (rule PrimRec1'.un_rec)
moreover have "UnaryRecOp g (pr_conv_3_to_1 h) = pr_conv_2_to_1 (PrimRecOp g h)" by (simp add: UnaryRecOp_def)
ultimately show "pr_conv_2_to_1 (PrimRecOp g h) ∈ PrimRec1'" by simp
qed
theorem pr_1_eq_1': "PrimRec1 = PrimRec1'"
proof -
have S1: "⋀ f. f ∈ PrimRec1 ⟶ f ∈ PrimRec1'" by (simp add: lm')
have S2: "⋀ f. f ∈ PrimRec1' ⟶ f ∈ PrimRec1" by (simp add: primrec'_into_primrec)
from S1 S2 show ?thesis by blast
qed
subsection ‹The scheme datatype›