theory Nat imports Identity begin axiomatization Nat :: \o\ and zero :: \o\ ("0") and suc :: \o \ o\ and NatInd :: \(o \ o) \ o \ (o \ o \ o) \ o \ o\ where NatF: "Nat: U i" and Nat_zero: "0: Nat" and Nat_suc: "n: Nat \ suc n: Nat" and NatE: "\ n: Nat; c\<^sub>0: C 0; \k rec. \k: Nat; rec: C k\ \ f k rec: C (suc k); \n. n: Nat \ C n: U i \ \ NatInd (fn n. C n) c\<^sub>0 (fn k rec. f k rec) n: C n" and Nat_comp_zero: "\ c\<^sub>0: C 0; \k rec. \k: Nat; rec: C k\ \ f k rec: C (suc k); \n. n: Nat \ C n: U i \ \ NatInd (fn n. C n) c\<^sub>0 (fn k rec. f k rec) 0 \ c\<^sub>0" and Nat_comp_suc: "\ n: Nat; c\<^sub>0: C 0; \k rec. \k: Nat; rec: C k\ \ f k rec: C (suc k); \n. n: Nat \ C n: U i \ \ NatInd (fn n. C n) c\<^sub>0 (fn k rec. f k rec) (suc n) \ f n (NatInd (fn n. C n) c\<^sub>0 (fn k rec. f k rec) n)" lemmas [form] = NatF and [intr, intro] = Nat_zero Nat_suc and [elim "?n"] = NatE and [comp] = Nat_comp_zero Nat_comp_suc abbreviation "NatRec C \ NatInd (fn _. C)" abbreviation one ("1") where "1 \ suc 0" abbreviation two ("2") where "2 \ suc 1" abbreviation three ("3") where "3 \ suc 2" abbreviation four ("4") where "4 \ suc 3" abbreviation five ("5") where "5 \ suc 4" abbreviation six ("6") where "6 \ suc 5" abbreviation seven ("7") where "7 \ suc 6" abbreviation eight ("8") where "8 \ suc 7" abbreviation nine ("9") where "9 \ suc 8" section \Basic arithmetic\ subsection \Addition\ definition add (infixl "+" 120) where "m + n \ NatRec Nat m (K suc) n" Lemma add_type [type]: assumes "m: Nat" "n: Nat" shows "m + n: Nat" unfolding add_def by typechk Lemma add_zero [comp]: assumes "m: Nat" shows "m + 0 \ m" unfolding add_def by compute Lemma add_suc [comp]: assumes "m: Nat" "n: Nat" shows "m + suc n \ suc (m + n)" unfolding add_def by compute Lemma (def) zero_add: assumes "n: Nat" shows "0 + n = n" apply (elim n) \<^item> by (compute; intro) \<^item> vars _ ih by compute (eq ih; refl) done Lemma (def) suc_add: assumes "m: Nat" "n: Nat" shows "suc m + n = suc (m + n)" apply (elim n) \<^item> by compute refl \<^item> vars _ ih by compute (eq ih; refl) done Lemma (def) suc_eq: assumes "m: Nat" "n: Nat" shows "p: m = n \ suc m = suc n" by (eq p) intro Lemma (def) add_assoc: assumes "l: Nat" "m: Nat" "n: Nat" shows "l + (m + n) = l + m+ n" apply (elim n) \<^item> by compute intro \<^item> vars _ ih by compute (eq ih; refl) done Lemma (def) add_comm: assumes "m: Nat" "n: Nat" shows "m + n = n + m" apply (elim n) \<^item> by (compute; rule zero_add[symmetric]) \<^item> vars n ih proof compute have "suc (m + n) = suc (n + m)" by (eq ih) refl also have ".. = suc n + m" by (rewr eq: suc_add) refl finally show "?" by infer qed done subsection \Multiplication\ definition mul (infixl "*" 121) where "m * n \ NatRec Nat 0 (K $ add m) n" Lemma mul_type [type]: assumes "m: Nat" "n: Nat" shows "m * n: Nat" unfolding mul_def by typechk Lemma mul_zero [comp]: assumes "n: Nat" shows "n * 0 \ 0" unfolding mul_def by compute Lemma mul_one [comp]: assumes "n: Nat" shows "n * 1 \ n" unfolding mul_def by compute Lemma mul_suc [comp]: assumes "m: Nat" "n: Nat" shows "m * suc n \ m + m * n" unfolding mul_def by compute Lemma (def) zero_mul: assumes "n: Nat" shows "0 * n = 0" apply (elim n) \<^item> by compute refl \<^item> vars n ih proof compute have "0 + 0 * n = 0 + 0 " by (eq ih) refl also have ".. = 0" by compute refl finally show "?" by infer qed done Lemma (def) suc_mul: assumes "m: Nat" "n: Nat" shows "suc m * n = m * n + n" apply (elim n) \<^item> by compute refl \<^item> vars n ih proof (compute, rewr eq: \ih:_\) have "suc m + (m * n + n) = suc (m + ?)" by (rule suc_add) also have ".. = suc (m + m * n + n)" by (rewr eq: add_assoc) refl finally show "?" by infer qed done Lemma (def) mul_dist_add: assumes "l: Nat" "m: Nat" "n: Nat" shows "l * (m + n) = l * m + l * n" apply (elim n) \<^item> by compute refl \<^item> vars n ih proof compute have "l + l * (m + n) = l + (l * m + l * n)" by (eq ih) refl also have ".. = l + l * m + l * n" by (rule add_assoc) also have ".. = l * m + l + l * n" by (rewr eq: add_comm) refl also have ".. = l * m + (l + l * n)" by (rewr eq: add_assoc) refl finally show "?" by infer qed done Lemma (def) mul_assoc: assumes "l: Nat" "m: Nat" "n: Nat" shows "l * (m * n) = l * m * n" apply (elim n) \<^item> by compute refl \<^item> vars n ih proof compute have "l * (m + m * n) = l * m + l * (m * n)" by (rule mul_dist_add) also have ".. = l * m + l * m * n" by (rewr eq: \ih:_\) refl finally show "?" by infer qed done Lemma (def) mul_comm: assumes "m: Nat" "n: Nat" shows "m * n = n * m" apply (elim n) \<^item> by compute (rewr eq: zero_mul, refl) \<^item> vars n ih proof (compute, rule pathinv) have "suc n * m = n * m + m" by (rule suc_mul) also have ".. = m + m * n" by (rewr eq: \ih:_\, rewr eq: add_comm) refl finally show "?" by infer qed done end