From 720da0f918118388d114e09664b129d2b29be2b1 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Fri, 22 May 2020 15:43:14 +0200 Subject: some tweaks and comments, in preparation for general inductive type elimination --- hott/Nat.thy | 51 +++++++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 47 insertions(+), 4 deletions(-) (limited to 'hott/Nat.thy') diff --git a/hott/Nat.thy b/hott/Nat.thy index 59ec517..3d938cd 100644 --- a/hott/Nat.thy +++ b/hott/Nat.thy @@ -49,10 +49,53 @@ abbreviation "NatRec C \ NatInd (K C)" section \Basic arithmetic\ -definition add (infixl "+" 50) where - [comps]: "m + n \ NatRec Nat n (K suc) m" - -definition mul (infixl "*" 55) where +definition add (infixl "+" 120) where "m + n \ NatRec Nat n (K suc) m" + +lemma add_type [typechk]: + assumes "m: Nat" "n: Nat" + shows "m + n: Nat" + unfolding add_def by typechk + +lemma add_rec [comps]: + assumes "m: Nat" "n: Nat" + shows "suc m + n \ suc (m + n)" + unfolding add_def by reduce + +lemma zero_add [comps]: + assumes "n: Nat" + shows "0 + n \ n" + unfolding add_def by reduce + +Lemma (derive) add_zero: + assumes "n: Nat" + shows "n + 0 = n" + apply (elim n) + \ by (reduce; intro) + \ vars _ IH by reduce (elim IH; intro) + done + +Lemma (derive) add_assoc: + assumes "l: Nat" "m: Nat" "n: Nat" + shows "l + (m + n) = (l + m) + n" + apply (elim l) + \ by reduce intro + \ vars _ IH by reduce (elim IH; intro) + done + +Lemma (derive) add_comm: + assumes "m: Nat" "n: Nat" + shows "m + n = n + m" + apply (elim m) + \ by (reduce; rule add_zero[sym]) + \ vars m p + apply (elim n) + (*Should also change the `n`s in the premises!*) + (*FIXME: Elimination tactic needs to do the same kind of thing as the + equality tactic with pushing context premises into the statement, applying + the appropriate elimination rule and then pulling back out.*) +oops + +definition mul (infixl "*" 120) where [comps]: "m * n \ NatRec Nat 0 (K $ add n) m" -- cgit v1.2.3