(* Title: HoTT/Nat.thy Author: Josh Chen Natural numbers *) theory Nat imports HoTT_Base begin section \Constants and type rules\ axiomatization Nat :: t ("\") and zero :: t ("0") and succ :: "t \ t" and indNat :: "[[t, t] \ t, t, t] \ t" ("(1ind\<^sub>\)") where Nat_form: "\: U O" and Nat_intro_0: "0: \" and Nat_intro_succ: "n: \ \ succ n: \" and Nat_elim: "\ C: \ \ U i; \n c. \n: \; c: C n\ \ f n c: C (succ n); a: C 0; n: \ \ \ ind\<^sub>\ f a n: C n" and Nat_comp_0: "\ C: \ \ U i; \n c. \n: \; c: C(n)\ \ f n c: C (succ n); a: C 0 \ \ ind\<^sub>\ f a 0 \ a" and Nat_comp_succ: "\ C: \ \ U i; \n c. \n: \; c: C n\ \ f n c: C (succ n); a: C 0; n: \ \ \ ind\<^sub>\ f a (succ n) \ f n (ind\<^sub>\ f a n)" text "Rule attribute declarations:" lemmas Nat_intro = Nat_intro_0 Nat_intro_succ lemmas Nat_comp [comp] = Nat_comp_0 Nat_comp_succ lemmas Nat_routine [intro] = Nat_form Nat_intro Nat_elim end