theory Prelude imports MLTT begin section \Sum type\ axiomatization Sum :: \o \ o \ o\ and inl :: \o \ o \ o \ o\ and inr :: \o \ o \ o \ o\ and SumInd :: \o \ o \ (o \ o) \ (o \ o) \ (o \ o) \ o \ o\ notation Sum (infixl "\" 50) axiomatization where SumF: "\A: U i; B: U i\ \ A \ B: U i" and Sum_inl: "\B: U i; a: A\ \ inl A B a: A \ B" and Sum_inr: "\A: U i; b: B\ \ inr A B b: A \ B" and SumE: "\ s: A \ B; \s. s: A \ B \ C s: U i; \a. a: A \ c a: C (inl A B a); \b. b: B \ d b: C (inr A B b) \ \ SumInd A B (fn s. C s) (fn a. c a) (fn b. d b) s: C s" and Sum_comp_inl: "\ a: A; \s. s: A \ B \ C s: U i; \a. a: A \ c a: C (inl A B a); \b. b: B \ d b: C (inr A B b) \ \ SumInd A B (fn s. C s) (fn a. c a) (fn b. d b) (inl A B a) \ c a" and Sum_comp_inr: "\ b: B; \s. s: A \ B \ C s: U i; \a. a: A \ c a: C (inl A B a); \b. b: B \ d b: C (inr A B b) \ \ SumInd A B (fn s. C s) (fn a. c a) (fn b. d b) (inr A B b) \ d b" lemmas [form] = SumF and [intr] = Sum_inl Sum_inr and [intro] = Sum_inl[rotated] Sum_inr[rotated] and [elim ?s] = SumE and [comp] = Sum_comp_inl Sum_comp_inr method left = rule Sum_inl method right = rule Sum_inr section \Empty and unit types\ axiomatization Top :: \o\ and tt :: \o\ and TopInd :: \(o \ o) \ o \ o \ o\ and Bot :: \o\ and BotInd :: \(o \ o) \ o \ o\ notation Top ("\") and Bot ("\") axiomatization where TopF: "\: U i" and TopI: "tt: \" and TopE: "\a: \; \x. x: \ \ C x: U i; c: C tt\ \ TopInd (fn x. C x) c a: C a" and Top_comp: "\\x. x: \ \ C x: U i; c: C tt\ \ TopInd (fn x. C x) c tt \ c" and BotF: "\: U i" and BotE: "\x: \; \x. x: \ \ C x: U i\ \ BotInd (fn x. C x) x: C x" lemmas [form] = TopF BotF and [intr, intro] = TopI and [elim ?a] = TopE and [elim ?x] = BotE and [comp] = Top_comp abbreviation (input) Not ("\_" [1000] 1000) where "\A \ A \ \" section \Booleans\ definition "Bool \ \ \ \" definition "true \ inl \ \ tt" definition "false \ inr \ \ tt" Lemma BoolF: "Bool: U i" and Bool_true: "true: Bool" and Bool_false: "false: Bool" unfolding Bool_def true_def false_def by typechk+ \ \Definitions like these should be handled by a future function package\ Lemma (def) ifelse [rotated 1]: assumes *[unfolded Bool_def true_def false_def]: "\x. x: Bool \ C x: U i" "x: Bool" "a: C true" "b: C false" shows "C x" using assms[unfolded Bool_def true_def false_def, type] by (elim x) (elim, fact)+ Lemma if_true: assumes "\x. x: Bool \ C x: U i" "a: C true" "b: C false" shows "ifelse C true a b \ a" unfolding ifelse_def true_def using assms unfolding Bool_def true_def false_def by compute Lemma if_false: assumes "\x. x: Bool \ C x: U i" "a: C true" "b: C false" shows "ifelse C false a b \ b" unfolding ifelse_def false_def using assms unfolding Bool_def true_def false_def by compute lemmas [form] = BoolF and [intr, intro] = Bool_true Bool_false and [comp] = if_true if_false and [elim ?x] = ifelse lemmas BoolE = ifelse subsection \Notation\ definition ifelse_i ("if _ then _ else _") where [implicit]: "if x then a else b \ ifelse {} x a b" translations "if x then a else b" \ "CONST ifelse C x a b" subsection \Logical connectives\ definition not ("!_") where "!x \ ifelse (K Bool) x false true" end