# HG changeset patch # User wenzelm # Date 1468588744 -7200 # Node ID 42e1dece537af5dd09e2aa47ec72324bf6bba336 # Parent 7f0e36eb73b4e956007e2c199353cc306eb55ee0 misc tuning and modernization; proper document setup; diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/Arith.thy --- a/src/CTT/Arith.thy Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/Arith.thy Fri Jul 15 15:19:04 2016 +0200 @@ -6,141 +6,115 @@ section \Elementary arithmetic\ theory Arith -imports Bool + imports Bool begin subsection \Arithmetic operators and their definitions\ -definition - add :: "[i,i]\i" (infixr "#+" 65) where - "a#+b \ rec(a, b, \u v. succ(v))" +definition add :: "[i,i]\i" (infixr "#+" 65) + where "a#+b \ rec(a, b, \u v. succ(v))" -definition - diff :: "[i,i]\i" (infixr "-" 65) where - "a-b \ rec(b, a, \u v. rec(v, 0, \x y. x))" +definition diff :: "[i,i]\i" (infixr "-" 65) + where "a-b \ rec(b, a, \u v. rec(v, 0, \x y. x))" -definition - absdiff :: "[i,i]\i" (infixr "|-|" 65) where - "a|-|b \ (a-b) #+ (b-a)" +definition absdiff :: "[i,i]\i" (infixr "|-|" 65) + where "a|-|b \ (a-b) #+ (b-a)" -definition - mult :: "[i,i]\i" (infixr "#*" 70) where - "a#*b \ rec(a, 0, \u v. b #+ v)" +definition mult :: "[i,i]\i" (infixr "#*" 70) + where "a#*b \ rec(a, 0, \u v. b #+ v)" -definition - mod :: "[i,i]\i" (infixr "mod" 70) where - "a mod b \ rec(a, 0, \u v. rec(succ(v) |-| b, 0, \x y. succ(v)))" +definition mod :: "[i,i]\i" (infixr "mod" 70) + where "a mod b \ rec(a, 0, \u v. rec(succ(v) |-| b, 0, \x y. succ(v)))" -definition - div :: "[i,i]\i" (infixr "div" 70) where - "a div b \ rec(a, 0, \u v. rec(succ(u) mod b, succ(v), \x y. v))" - +definition div :: "[i,i]\i" (infixr "div" 70) + where "a div b \ rec(a, 0, \u v. rec(succ(u) mod b, succ(v), \x y. v))" lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def subsection \Proofs about elementary arithmetic: addition, multiplication, etc.\ -(** Addition *) +subsubsection \Addition\ -(*typing of add: short and long versions*) +text \Typing of \add\: short and long versions.\ lemma add_typing: "\a:N; b:N\ \ a #+ b : N" -apply (unfold arith_defs) -apply typechk -done + unfolding arith_defs by typechk lemma add_typingL: "\a = c:N; b = d:N\ \ a #+ b = c #+ d : N" -apply (unfold arith_defs) -apply equal -done + unfolding arith_defs by equal -(*computation for add: 0 and successor cases*) +text \Computation for \add\: 0 and successor cases.\ lemma addC0: "b:N \ 0 #+ b = b : N" -apply (unfold arith_defs) -apply rew -done + unfolding arith_defs by rew lemma addC_succ: "\a:N; b:N\ \ succ(a) #+ b = succ(a #+ b) : N" -apply (unfold arith_defs) -apply rew -done + unfolding arith_defs by rew -(** Multiplication *) +subsubsection \Multiplication\ -(*typing of mult: short and long versions*) +text \Typing of \mult\: short and long versions.\ lemma mult_typing: "\a:N; b:N\ \ a #* b : N" -apply (unfold arith_defs) -apply (typechk add_typing) -done + unfolding arith_defs by (typechk add_typing) lemma mult_typingL: "\a = c:N; b = d:N\ \ a #* b = c #* d : N" -apply (unfold arith_defs) -apply (equal add_typingL) -done + unfolding arith_defs by (equal add_typingL) -(*computation for mult: 0 and successor cases*) + +text \Computation for \mult\: 0 and successor cases.\ lemma multC0: "b:N \ 0 #* b = 0 : N" -apply (unfold arith_defs) -apply rew -done + unfolding arith_defs by rew lemma multC_succ: "\a:N; b:N\ \ succ(a) #* b = b #+ (a #* b) : N" -apply (unfold arith_defs) -apply rew -done + unfolding arith_defs by rew -(** Difference *) +subsubsection \Difference\ -(*typing of difference*) +text \Typing of difference.\ lemma diff_typing: "\a:N; b:N\ \ a - b : N" -apply (unfold arith_defs) -apply typechk -done + unfolding arith_defs by typechk lemma diff_typingL: "\a = c:N; b = d:N\ \ a - b = c - d : N" -apply (unfold arith_defs) -apply equal -done + unfolding arith_defs by equal -(*computation for difference: 0 and successor cases*) +text \Computation for difference: 0 and successor cases.\ lemma diffC0: "a:N \ a - 0 = a : N" -apply (unfold arith_defs) -apply rew -done + unfolding arith_defs by rew -(*Note: rec(a, 0, \z w.z) is pred(a). *) +text \Note: \rec(a, 0, \z w.z)\ is \pred(a).\\ lemma diff_0_eq_0: "b:N \ 0 - b = 0 : N" -apply (unfold arith_defs) -apply (NE b) -apply hyp_rew -done - + unfolding arith_defs + apply (NE b) + apply hyp_rew + done -(*Essential to simplify FIRST!! (Else we get a critical pair) - succ(a) - succ(b) rewrites to pred(succ(a) - b) *) +text \ + Essential to simplify FIRST!! (Else we get a critical pair) + \succ(a) - succ(b)\ rewrites to \pred(succ(a) - b)\. +\ lemma diff_succ_succ: "\a:N; b:N\ \ succ(a) - succ(b) = a - b : N" -apply (unfold arith_defs) -apply hyp_rew -apply (NE b) -apply hyp_rew -done + unfolding arith_defs + apply hyp_rew + apply (NE b) + apply hyp_rew + done subsection \Simplification\ lemmas arith_typing_rls = add_typing mult_typing diff_typing and arith_congr_rls = add_typingL mult_typingL diff_typingL + lemmas congr_rls = arith_congr_rls intrL2_rls elimL_rls lemmas arithC_rls = @@ -149,30 +123,23 @@ diffC0 diff_0_eq_0 diff_succ_succ ML \ + structure Arith_simp = TSimpFun( + val refl = @{thm refl_elem} + val sym = @{thm sym_elem} + val trans = @{thm trans_elem} + val refl_red = @{thm refl_red} + val trans_red = @{thm trans_red} + val red_if_equal = @{thm red_if_equal} + val default_rls = @{thms arithC_rls comp_rls} + val routine_tac = routine_tac @{thms arith_typing_rls routine_rls} + ) -structure Arith_simp_data: TSIMP_DATA = - struct - val refl = @{thm refl_elem} - val sym = @{thm sym_elem} - val trans = @{thm trans_elem} - val refl_red = @{thm refl_red} - val trans_red = @{thm trans_red} - val red_if_equal = @{thm red_if_equal} - val default_rls = @{thms arithC_rls} @ @{thms comp_rls} - val routine_tac = routine_tac (@{thms arith_typing_rls} @ @{thms routine_rls}) - end + fun arith_rew_tac ctxt prems = + make_rew_tac ctxt (Arith_simp.norm_tac ctxt (@{thms congr_rls}, prems)) -structure Arith_simp = TSimpFun (Arith_simp_data) - -local val congr_rls = @{thms congr_rls} in - -fun arith_rew_tac ctxt prems = make_rew_tac ctxt - (Arith_simp.norm_tac ctxt (congr_rls, prems)) - -fun hyp_arith_rew_tac ctxt prems = make_rew_tac ctxt - (Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, congr_rls, prems)) - -end + fun hyp_arith_rew_tac ctxt prems = + make_rew_tac ctxt + (Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, @{thms congr_rls}, prems)) \ method_setup arith_rew = \ @@ -186,284 +153,263 @@ subsection \Addition\ -(*Associative law for addition*) +text \Associative law for addition.\ lemma add_assoc: "\a:N; b:N; c:N\ \ (a #+ b) #+ c = a #+ (b #+ c) : N" -apply (NE a) -apply hyp_arith_rew -done + apply (NE a) + apply hyp_arith_rew + done - -(*Commutative law for addition. Can be proved using three inductions. - Must simplify after first induction! Orientation of rewrites is delicate*) +text \Commutative law for addition. Can be proved using three inductions. + Must simplify after first induction! Orientation of rewrites is delicate.\ lemma add_commute: "\a:N; b:N\ \ a #+ b = b #+ a : N" -apply (NE a) -apply hyp_arith_rew -apply (rule sym_elem) -prefer 2 -apply (NE b) -prefer 4 -apply (NE b) -apply hyp_arith_rew -done + apply (NE a) + apply hyp_arith_rew + apply (rule sym_elem) + prefer 2 + apply (NE b) + prefer 4 + apply (NE b) + apply hyp_arith_rew + done subsection \Multiplication\ -(*right annihilation in product*) +text \Right annihilation in product.\ lemma mult_0_right: "a:N \ a #* 0 = 0 : N" -apply (NE a) -apply hyp_arith_rew -done + apply (NE a) + apply hyp_arith_rew + done -(*right successor law for multiplication*) +text \Right successor law for multiplication.\ lemma mult_succ_right: "\a:N; b:N\ \ a #* succ(b) = a #+ (a #* b) : N" -apply (NE a) -apply (hyp_arith_rew add_assoc [THEN sym_elem]) -apply (assumption | rule add_commute mult_typingL add_typingL intrL_rls refl_elem)+ -done + apply (NE a) + apply (hyp_arith_rew add_assoc [THEN sym_elem]) + apply (assumption | rule add_commute mult_typingL add_typingL intrL_rls refl_elem)+ + done -(*Commutative law for multiplication*) +text \Commutative law for multiplication.\ lemma mult_commute: "\a:N; b:N\ \ a #* b = b #* a : N" -apply (NE a) -apply (hyp_arith_rew mult_0_right mult_succ_right) -done + apply (NE a) + apply (hyp_arith_rew mult_0_right mult_succ_right) + done -(*addition distributes over multiplication*) +text \Addition distributes over multiplication.\ lemma add_mult_distrib: "\a:N; b:N; c:N\ \ (a #+ b) #* c = (a #* c) #+ (b #* c) : N" -apply (NE a) -apply (hyp_arith_rew add_assoc [THEN sym_elem]) -done + apply (NE a) + apply (hyp_arith_rew add_assoc [THEN sym_elem]) + done -(*Associative law for multiplication*) +text \Associative law for multiplication.\ lemma mult_assoc: "\a:N; b:N; c:N\ \ (a #* b) #* c = a #* (b #* c) : N" -apply (NE a) -apply (hyp_arith_rew add_mult_distrib) -done + apply (NE a) + apply (hyp_arith_rew add_mult_distrib) + done subsection \Difference\ text \ -Difference on natural numbers, without negative numbers - a - b = 0 iff a<=b a - b = succ(c) iff a>b\ + Difference on natural numbers, without negative numbers + \<^item> \a - b = 0\ iff \a \ b\ + \<^item> \a - b = succ(c)\ iff \a > b\ +\ lemma diff_self_eq_0: "a:N \ a - a = 0 : N" -apply (NE a) -apply hyp_arith_rew -done + apply (NE a) + apply hyp_arith_rew + done lemma add_0_right: "\c : N; 0 : N; c : N\ \ c #+ 0 = c : N" by (rule addC0 [THEN [3] add_commute [THEN trans_elem]]) -(*Addition is the inverse of subtraction: if b<=x then b#+(x-b) = x. +text \ + Addition is the inverse of subtraction: if \b \ x\ then \b #+ (x - b) = x\. An example of induction over a quantified formula (a product). - Uses rewriting with a quantified, implicative inductive hypothesis.*) + Uses rewriting with a quantified, implicative inductive hypothesis. +\ schematic_goal add_diff_inverse_lemma: "b:N \ ?a : \x:N. Eq(N, b-x, 0) \ Eq(N, b #+ (x-b), x)" -apply (NE b) -(*strip one "universal quantifier" but not the "implication"*) -apply (rule_tac [3] intr_rls) -(*case analysis on x in - (succ(u) <= x) \ (succ(u)#+(x-succ(u)) = x) *) -prefer 4 -apply (NE x) -apply assumption -(*Prepare for simplification of types -- the antecedent succ(u)<=x *) -apply (rule_tac [2] replace_type) -apply (rule_tac [1] replace_type) -apply arith_rew -(*Solves first 0 goal, simplifies others. Two sugbgoals remain. - Both follow by rewriting, (2) using quantified induction hyp*) -apply intr (*strips remaining PRODs*) -apply (hyp_arith_rew add_0_right) -apply assumption -done + apply (NE b) + \ \strip one "universal quantifier" but not the "implication"\ + apply (rule_tac [3] intr_rls) + \ \case analysis on \x\ in \succ(u) \ x \ succ(u) #+ (x - succ(u)) = x\\ + prefer 4 + apply (NE x) + apply assumption + \ \Prepare for simplification of types -- the antecedent \succ(u) \ x\\ + apply (rule_tac [2] replace_type) + apply (rule_tac [1] replace_type) + apply arith_rew + \ \Solves first 0 goal, simplifies others. Two sugbgoals remain. + Both follow by rewriting, (2) using quantified induction hyp.\ + apply intr \ \strips remaining \\\s\ + apply (hyp_arith_rew add_0_right) + apply assumption + done - -(*Version of above with premise b-a=0 i.e. a >= b. - Using ProdE does not work -- for ?B(?a) is ambiguous. - Instead, add_diff_inverse_lemma states the desired induction scheme - the use of RS below instantiates Vars in ProdE automatically. *) +text \ + Version of above with premise \b - a = 0\ i.e. \a \ b\. + Using @{thm ProdE} does not work -- for \?B(?a)\ is ambiguous. + Instead, @{thm add_diff_inverse_lemma} states the desired induction scheme; + the use of \THEN\ below instantiates Vars in @{thm ProdE} automatically. +\ lemma add_diff_inverse: "\a:N; b:N; b - a = 0 : N\ \ b #+ (a-b) = a : N" -apply (rule EqE) -apply (rule add_diff_inverse_lemma [THEN ProdE, THEN ProdE]) -apply (assumption | rule EqI)+ -done + apply (rule EqE) + apply (rule add_diff_inverse_lemma [THEN ProdE, THEN ProdE]) + apply (assumption | rule EqI)+ + done subsection \Absolute difference\ -(*typing of absolute difference: short and long versions*) +text \Typing of absolute difference: short and long versions.\ lemma absdiff_typing: "\a:N; b:N\ \ a |-| b : N" -apply (unfold arith_defs) -apply typechk -done + unfolding arith_defs by typechk lemma absdiff_typingL: "\a = c:N; b = d:N\ \ a |-| b = c |-| d : N" -apply (unfold arith_defs) -apply equal -done + unfolding arith_defs by equal lemma absdiff_self_eq_0: "a:N \ a |-| a = 0 : N" -apply (unfold absdiff_def) -apply (arith_rew diff_self_eq_0) -done + unfolding absdiff_def by (arith_rew diff_self_eq_0) lemma absdiffC0: "a:N \ 0 |-| a = a : N" -apply (unfold absdiff_def) -apply hyp_arith_rew -done - + unfolding absdiff_def by hyp_arith_rew lemma absdiff_succ_succ: "\a:N; b:N\ \ succ(a) |-| succ(b) = a |-| b : N" -apply (unfold absdiff_def) -apply hyp_arith_rew -done + unfolding absdiff_def by hyp_arith_rew -(*Note how easy using commutative laws can be? ...not always... *) +text \Note how easy using commutative laws can be? ...not always...\ lemma absdiff_commute: "\a:N; b:N\ \ a |-| b = b |-| a : N" -apply (unfold absdiff_def) -apply (rule add_commute) -apply (typechk diff_typing) -done + unfolding absdiff_def + apply (rule add_commute) + apply (typechk diff_typing) + done -(*If a+b=0 then a=0. Surprisingly tedious*) +text \If \a + b = 0\ then \a = 0\. Surprisingly tedious.\ schematic_goal add_eq0_lemma: "\a:N; b:N\ \ ?c : \u: Eq(N,a#+b,0) . Eq(N,a,0)" -apply (NE a) -apply (rule_tac [3] replace_type) -apply arith_rew -apply intr (*strips remaining PRODs*) -apply (rule_tac [2] zero_ne_succ [THEN FE]) -apply (erule_tac [3] EqE [THEN sym_elem]) -apply (typechk add_typing) -done + apply (NE a) + apply (rule_tac [3] replace_type) + apply arith_rew + apply intr \ \strips remaining \\\s\ + apply (rule_tac [2] zero_ne_succ [THEN FE]) + apply (erule_tac [3] EqE [THEN sym_elem]) + apply (typechk add_typing) + done -(*Version of above with the premise a+b=0. - Again, resolution instantiates variables in ProdE *) +text \ + Version of above with the premise \a + b = 0\. + Again, resolution instantiates variables in @{thm ProdE}. +\ lemma add_eq0: "\a:N; b:N; a #+ b = 0 : N\ \ a = 0 : N" -apply (rule EqE) -apply (rule add_eq0_lemma [THEN ProdE]) -apply (rule_tac [3] EqI) -apply typechk -done + apply (rule EqE) + apply (rule add_eq0_lemma [THEN ProdE]) + apply (rule_tac [3] EqI) + apply typechk + done -(*Here is a lemma to infer a-b=0 and b-a=0 from a|-|b=0, below. *) +text \Here is a lemma to infer \a - b = 0\ and \b - a = 0\ from \a |-| b = 0\, below.\ schematic_goal absdiff_eq0_lem: "\a:N; b:N; a |-| b = 0 : N\ \ ?a : \v: Eq(N, a-b, 0) . Eq(N, b-a, 0)" -apply (unfold absdiff_def) -apply intr -apply eqintr -apply (rule_tac [2] add_eq0) -apply (rule add_eq0) -apply (rule_tac [6] add_commute [THEN trans_elem]) -apply (typechk diff_typing) -done + apply (unfold absdiff_def) + apply intr + apply eqintr + apply (rule_tac [2] add_eq0) + apply (rule add_eq0) + apply (rule_tac [6] add_commute [THEN trans_elem]) + apply (typechk diff_typing) + done -(*if a |-| b = 0 then a = b - proof: a-b=0 and b-a=0, so b = a+(b-a) = a+0 = a*) +text \If \a |-| b = 0\ then \a = b\ + proof: \a - b = 0\ and \b - a = 0\, so \b = a + (b - a) = a + 0 = a\. +\ lemma absdiff_eq0: "\a |-| b = 0 : N; a:N; b:N\ \ a = b : N" -apply (rule EqE) -apply (rule absdiff_eq0_lem [THEN SumE]) -apply eqintr -apply (rule add_diff_inverse [THEN sym_elem, THEN trans_elem]) -apply (erule_tac [3] EqE) -apply (hyp_arith_rew add_0_right) -done + apply (rule EqE) + apply (rule absdiff_eq0_lem [THEN SumE]) + apply eqintr + apply (rule add_diff_inverse [THEN sym_elem, THEN trans_elem]) + apply (erule_tac [3] EqE) + apply (hyp_arith_rew add_0_right) + done subsection \Remainder and Quotient\ -(*typing of remainder: short and long versions*) +text \Typing of remainder: short and long versions.\ lemma mod_typing: "\a:N; b:N\ \ a mod b : N" -apply (unfold mod_def) -apply (typechk absdiff_typing) -done + unfolding mod_def by (typechk absdiff_typing) lemma mod_typingL: "\a = c:N; b = d:N\ \ a mod b = c mod d : N" -apply (unfold mod_def) -apply (equal absdiff_typingL) -done + unfolding mod_def by (equal absdiff_typingL) -(*computation for mod : 0 and successor cases*) +text \Computation for \mod\: 0 and successor cases.\ lemma modC0: "b:N \ 0 mod b = 0 : N" -apply (unfold mod_def) -apply (rew absdiff_typing) -done + unfolding mod_def by (rew absdiff_typing) lemma modC_succ: "\a:N; b:N\ \ succ(a) mod b = rec(succ(a mod b) |-| b, 0, \x y. succ(a mod b)) : N" -apply (unfold mod_def) -apply (rew absdiff_typing) -done + unfolding mod_def by (rew absdiff_typing) -(*typing of quotient: short and long versions*) +text \Typing of quotient: short and long versions.\ lemma div_typing: "\a:N; b:N\ \ a div b : N" -apply (unfold div_def) -apply (typechk absdiff_typing mod_typing) -done + unfolding div_def by (typechk absdiff_typing mod_typing) lemma div_typingL: "\a = c:N; b = d:N\ \ a div b = c div d : N" -apply (unfold div_def) -apply (equal absdiff_typingL mod_typingL) -done + unfolding div_def by (equal absdiff_typingL mod_typingL) lemmas div_typing_rls = mod_typing div_typing absdiff_typing -(*computation for quotient: 0 and successor cases*) +text \Computation for quotient: 0 and successor cases.\ lemma divC0: "b:N \ 0 div b = 0 : N" -apply (unfold div_def) -apply (rew mod_typing absdiff_typing) -done + unfolding div_def by (rew mod_typing absdiff_typing) lemma divC_succ: "\a:N; b:N\ \ succ(a) div b = rec(succ(a) mod b, succ(a div b), \x y. a div b) : N" -apply (unfold div_def) -apply (rew mod_typing) -done + unfolding div_def by (rew mod_typing) -(*Version of above with same condition as the mod one*) +text \Version of above with same condition as the \mod\ one.\ lemma divC_succ2: "\a:N; b:N\ \ succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), \x y. a div b) : N" -apply (rule divC_succ [THEN trans_elem]) -apply (rew div_typing_rls modC_succ) -apply (NE "succ (a mod b) |-|b") -apply (rew mod_typing div_typing absdiff_typing) -done + apply (rule divC_succ [THEN trans_elem]) + apply (rew div_typing_rls modC_succ) + apply (NE "succ (a mod b) |-|b") + apply (rew mod_typing div_typing absdiff_typing) + done -(*for case analysis on whether a number is 0 or a successor*) +text \For case analysis on whether a number is 0 or a successor.\ lemma iszero_decidable: "a:N \ rec(a, inl(eq), \ka kb. inr()) : - Eq(N,a,0) + (\x:N. Eq(N,a, succ(x)))" -apply (NE a) -apply (rule_tac [3] PlusI_inr) -apply (rule_tac [2] PlusI_inl) -apply eqintr -apply equal -done + Eq(N,a,0) + (\x:N. Eq(N,a, succ(x)))" + apply (NE a) + apply (rule_tac [3] PlusI_inr) + apply (rule_tac [2] PlusI_inl) + apply eqintr + apply equal + done -(*Main Result. Holds when b is 0 since a mod 0 = a and a div 0 = 0 *) +text \Main Result. Holds when \b\ is 0 since \a mod 0 = a\ and \a div 0 = 0\.\ lemma mod_div_equality: "\a:N; b:N\ \ a mod b #+ (a div b) #* b = a : N" -apply (NE a) -apply (arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2) -apply (rule EqE) -(*case analysis on succ(u mod b)|-|b *) -apply (rule_tac a1 = "succ (u mod b) |-| b" in iszero_decidable [THEN PlusE]) -apply (erule_tac [3] SumE) -apply (hyp_arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2) -(*Replace one occurrence of b by succ(u mod b). Clumsy!*) -apply (rule add_typingL [THEN trans_elem]) -apply (erule EqE [THEN absdiff_eq0, THEN sym_elem]) -apply (rule_tac [3] refl_elem) -apply (hyp_arith_rew div_typing_rls) -done + apply (NE a) + apply (arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2) + apply (rule EqE) + \ \case analysis on \succ(u mod b) |-| b\\ + apply (rule_tac a1 = "succ (u mod b) |-| b" in iszero_decidable [THEN PlusE]) + apply (erule_tac [3] SumE) + apply (hyp_arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2) + \ \Replace one occurrence of \b\ by \succ(u mod b)\. Clumsy!\ + apply (rule add_typingL [THEN trans_elem]) + apply (erule EqE [THEN absdiff_eq0, THEN sym_elem]) + apply (rule_tac [3] refl_elem) + apply (hyp_arith_rew div_typing_rls) + done end diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/Bool.thy --- a/src/CTT/Bool.thy Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/Bool.thy Fri Jul 15 15:19:04 2016 +0200 @@ -6,80 +6,68 @@ section \The two-element type (booleans and conditionals)\ theory Bool -imports CTT + imports CTT begin -definition - Bool :: "t" where - "Bool \ T+T" +definition Bool :: "t" + where "Bool \ T+T" -definition - true :: "i" where - "true \ inl(tt)" +definition true :: "i" + where "true \ inl(tt)" -definition - false :: "i" where - "false \ inr(tt)" +definition false :: "i" + where "false \ inr(tt)" -definition - cond :: "[i,i,i]\i" where - "cond(a,b,c) \ when(a, \u. b, \u. c)" +definition cond :: "[i,i,i]\i" + where "cond(a,b,c) \ when(a, \u. b, \u. c)" lemmas bool_defs = Bool_def true_def false_def cond_def -subsection \Derivation of rules for the type Bool\ +subsection \Derivation of rules for the type \Bool\\ -(*formation rule*) +text \Formation rule.\ lemma boolF: "Bool type" -apply (unfold bool_defs) -apply typechk -done + unfolding bool_defs by typechk - -(*introduction rules for true, false*) +text \Introduction rules for \true\, \false\.\ lemma boolI_true: "true : Bool" -apply (unfold bool_defs) -apply typechk -done + unfolding bool_defs by typechk lemma boolI_false: "false : Bool" -apply (unfold bool_defs) -apply typechk -done + unfolding bool_defs by typechk -(*elimination rule: typing of cond*) +text \Elimination rule: typing of \cond\.\ lemma boolE: "\p:Bool; a : C(true); b : C(false)\ \ cond(p,a,b) : C(p)" -apply (unfold bool_defs) -apply typechk -apply (erule_tac [!] TE) -apply typechk -done + unfolding bool_defs + apply (typechk; erule TE) + apply typechk + done lemma boolEL: "\p = q : Bool; a = c : C(true); b = d : C(false)\ \ cond(p,a,b) = cond(q,c,d) : C(p)" -apply (unfold bool_defs) -apply (rule PlusEL) -apply (erule asm_rl refl_elem [THEN TEL])+ -done + unfolding bool_defs + apply (rule PlusEL) + apply (erule asm_rl refl_elem [THEN TEL])+ + done -(*computation rules for true, false*) +text \Computation rules for \true\, \false\.\ lemma boolC_true: "\a : C(true); b : C(false)\ \ cond(true,a,b) = a : C(true)" -apply (unfold bool_defs) -apply (rule comp_rls) -apply typechk -apply (erule_tac [!] TE) -apply typechk -done + unfolding bool_defs + apply (rule comp_rls) + apply typechk + apply (erule_tac [!] TE) + apply typechk + done lemma boolC_false: "\a : C(true); b : C(false)\ \ cond(false,a,b) = b : C(false)" -apply (unfold bool_defs) -apply (rule comp_rls) -apply typechk -apply (erule_tac [!] TE) -apply typechk -done + unfolding bool_defs + apply (rule comp_rls) + apply typechk + apply (erule_tac [!] TE) + apply typechk + done end diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/CTT.thy --- a/src/CTT/CTT.thy Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/CTT.thy Fri Jul 15 15:19:04 2016 +0200 @@ -17,45 +17,46 @@ typedecl o consts - (*Types*) + \ \Types\ F :: "t" - T :: "t" (*F is empty, T contains one element*) + T :: "t" \ \\F\ is empty, \T\ contains one element\ contr :: "i\i" tt :: "i" - (*Natural numbers*) + \ \Natural numbers\ N :: "t" succ :: "i\i" rec :: "[i, i, [i,i]\i] \ i" - (*Unions*) + \ \Unions\ inl :: "i\i" inr :: "i\i" "when" :: "[i, i\i, i\i]\i" - (*General Sum and Binary Product*) + \ \General Sum and Binary Product\ Sum :: "[t, i\t]\t" fst :: "i\i" snd :: "i\i" split :: "[i, [i,i]\i] \i" - (*General Product and Function Space*) + \ \General Product and Function Space\ Prod :: "[t, i\t]\t" - (*Types*) + \ \Types\ Plus :: "[t,t]\t" (infixr "+" 40) - (*Equality type*) + \ \Equality type\ Eq :: "[t,i,i]\t" eq :: "i" - (*Judgements*) + \ \Judgements\ Type :: "t \ prop" ("(_ type)" [10] 5) Eqtype :: "[t,t]\prop" ("(_ =/ _)" [10,10] 5) Elem :: "[i, t]\prop" ("(_ /: _)" [10,10] 5) Eqelem :: "[i,i,t]\prop" ("(_ =/ _ :/ _)" [10,10,10] 5) Reduce :: "[i,i]\prop" ("Reduce[_,_]") - (*Types*) - (*Functions*) + \ \Types\ + + \ \Functions\ lambda :: "(i \ i) \ i" (binder "\<^bold>\" 10) app :: "[i,i]\i" (infixl "`" 60) - (*Natural numbers*) + \ \Natural numbers\ Zero :: "i" ("0") - (*Pairing*) + \ \Pairing\ pair :: "[i,i]\i" ("(1<_,/_>)") syntax @@ -65,35 +66,37 @@ "\x:A. B" \ "CONST Prod(A, \x. B)" "\x:A. B" \ "CONST Sum(A, \x. B)" -abbreviation - Arrow :: "[t,t]\t" (infixr "\" 30) where - "A \ B \ \_:A. B" -abbreviation - Times :: "[t,t]\t" (infixr "\" 50) where - "A \ B \ \_:A. B" +abbreviation Arrow :: "[t,t]\t" (infixr "\" 30) + where "A \ B \ \_:A. B" + +abbreviation Times :: "[t,t]\t" (infixr "\" 50) + where "A \ B \ \_:A. B" - (*Reduction: a weaker notion than equality; a hack for simplification. - Reduce[a,b] means either that a=b:A for some A or else that "a" and "b" - are textually identical.*) +text \ + Reduction: a weaker notion than equality; a hack for simplification. + \Reduce[a,b]\ means either that \a = b : A\ for some \A\ or else + that \a\ and \b\ are textually identical. - (*does not verify a:A! Sound because only trans_red uses a Reduce premise - No new theorems can be proved about the standard judgements.*) -axiomatization where + Does not verify \a:A\! Sound because only \trans_red\ uses a \Reduce\ + premise. No new theorems can be proved about the standard judgements. +\ +axiomatization +where refl_red: "\a. Reduce[a,a]" and red_if_equal: "\a b A. a = b : A \ Reduce[a,b]" and trans_red: "\a b c A. \a = b : A; Reduce[b,c]\ \ a = c : A" and - (*Reflexivity*) + \ \Reflexivity\ refl_type: "\A. A type \ A = A" and refl_elem: "\a A. a : A \ a = a : A" and - (*Symmetry*) + \ \Symmetry\ sym_type: "\A B. A = B \ B = A" and sym_elem: "\a b A. a = b : A \ b = a : A" and - (*Transitivity*) + \ \Transitivity\ trans_type: "\A B C. \A = B; B = C\ \ A = C" and trans_elem: "\a b c A. \a = b : A; b = c : A\ \ a = c : A" and @@ -101,7 +104,7 @@ equal_types: "\a A B. \a : A; A = B\ \ a : B" and equal_typesL: "\a b A B. \a = b : A; A = B\ \ a = b : B" and - (*Substitution*) + \ \Substitution\ subst_type: "\a A B. \a : A; \z. z:A \ B(z) type\ \ B(a) type" and subst_typeL: "\a c A B D. \a = c : A; \z. z:A \ B(z) = D(z)\ \ B(a) = D(c)" and @@ -111,7 +114,7 @@ "\a b c d A B. \a = c : A; \z. z:A \ b(z)=d(z) : B(z)\ \ b(a)=d(c) : B(a)" and - (*The type N -- natural numbers*) + \ \The type \N\ -- natural numbers\ NF: "N type" and NI0: "0 : N" and @@ -135,11 +138,11 @@ "\p a b C. \p: N; a: C(0); \u v. \u: N; v: C(u)\ \ b(u,v): C(succ(u))\ \ rec(succ(p), a, \u v. b(u,v)) = b(p, rec(p, a, \u v. b(u,v))) : C(succ(p))" and - (*The fourth Peano axiom. See page 91 of Martin-Löf's book*) + \ \The fourth Peano axiom. See page 91 of Martin-Löf's book.\ zero_ne_succ: "\a. \a: N; 0 = succ(a) : N\ \ 0: F" and - (*The Product of a family of types*) + \ \The Product of a family of types\ ProdF: "\A B. \A type; \x. x:A \ B(x) type\ \ \x:A. B(x) type" and @@ -160,7 +163,7 @@ ProdC2: "\p A B. p : \x:A. B(x) \ (\<^bold>\x. p`x) = p : \x:A. B(x)" and - (*The Sum of a family of types*) + \ \The Sum of a family of types\ SumF: "\A B. \A type; \x. x:A \ B(x) type\ \ \x:A. B(x) type" and SumFL: "\A B C D. \A = C; \x. x:A \ B(x) = D(x)\ \ \x:A. B(x) = \x:C. D(x)" and @@ -182,7 +185,7 @@ snd_def: "\a. snd(a) \ split(a, \x y. y)" and - (*The sum of two types*) + \ \The sum of two types\ PlusF: "\A B. \A type; B type\ \ A+B type" and PlusFL: "\A B C D. \A = C; B = D\ \ A+B = C+D" and @@ -217,27 +220,31 @@ \ when(inr(b), \x. c(x), \y. d(y)) = d(b) : C(inr(b))" and - (*The type Eq*) + \ \The type \Eq\\ EqF: "\a b A. \A type; a : A; b : A\ \ Eq(A,a,b) type" and EqFL: "\a b c d A B. \A = B; a = c : A; b = d : A\ \ Eq(A,a,b) = Eq(B,c,d)" and EqI: "\a b A. a = b : A \ eq : Eq(A,a,b)" and EqE: "\p a b A. p : Eq(A,a,b) \ a = b : A" and - (*By equality of types, can prove C(p) from C(eq), an elimination rule*) + \ \By equality of types, can prove \C(p)\ from \C(eq)\, an elimination rule\ EqC: "\p a b A. p : Eq(A,a,b) \ p = eq : Eq(A,a,b)" and - (*The type F*) + + \ \The type \F\\ FF: "F type" and FE: "\p C. \p: F; C type\ \ contr(p) : C" and FEL: "\p q C. \p = q : F; C type\ \ contr(p) = contr(q) : C" and - (*The type T - Martin-Löf's book (page 68) discusses elimination and computation. - Elimination can be derived by computation and equality of types, - but with an extra premise C(x) type x:T. - Also computation can be derived from elimination. *) + + \ \The type T\ + \ \ + Martin-Löf's book (page 68) discusses elimination and computation. + Elimination can be derived by computation and equality of types, + but with an extra premise \C(x)\ type \x:T\. + Also computation can be derived from elimination. + \ TF: "T type" and TI: "tt : T" and @@ -248,55 +255,59 @@ subsection "Tactics and derived rules for Constructive Type Theory" -(*Formation rules*) +text \Formation rules.\ lemmas form_rls = NF ProdF SumF PlusF EqF FF TF and formL_rls = ProdFL SumFL PlusFL EqFL -(*Introduction rules - OMITTED: EqI, because its premise is an eqelem, not an elem*) +text \ + Introduction rules. OMITTED: + \<^item> \EqI\, because its premise is an \eqelem\, not an \elem\. +\ lemmas intr_rls = NI0 NI_succ ProdI SumI PlusI_inl PlusI_inr TI and intrL_rls = NI_succL ProdIL SumIL PlusI_inlL PlusI_inrL -(*Elimination rules - OMITTED: EqE, because its conclusion is an eqelem, not an elem - TE, because it does not involve a constructor *) +text \ + Elimination rules. OMITTED: + \<^item> \EqE\, because its conclusion is an \eqelem\, not an \elem\ + \<^item> \TE\, because it does not involve a constructor. +\ lemmas elim_rls = NE ProdE SumE PlusE FE and elimL_rls = NEL ProdEL SumEL PlusEL FEL -(*OMITTED: eqC are TC because they make rewriting loop: p = un = un = ... *) +text \OMITTED: \eqC\ are \TC\ because they make rewriting loop: \p = un = un = \\\ lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr -(*rules with conclusion a:A, an elem judgement*) +text \Rules with conclusion \a:A\, an elem judgement.\ lemmas element_rls = intr_rls elim_rls -(*Definitions are (meta)equality axioms*) +text \Definitions are (meta)equality axioms.\ lemmas basic_defs = fst_def snd_def -(*Compare with standard version: B is applied to UNSIMPLIFIED expression! *) +text \Compare with standard version: \B\ is applied to UNSIMPLIFIED expression!\ lemma SumIL2: "\c = a : A; d = b : B(a)\ \ = : Sum(A,B)" -apply (rule sym_elem) -apply (rule SumIL) -apply (rule_tac [!] sym_elem) -apply assumption+ -done + apply (rule sym_elem) + apply (rule SumIL) + apply (rule_tac [!] sym_elem) + apply assumption+ + done lemmas intrL2_rls = NI_succL ProdIL SumIL2 PlusI_inlL PlusI_inrL -(*Exploit p:Prod(A,B) to create the assumption z:B(a). - A more natural form of product elimination. *) +text \ + Exploit \p:Prod(A,B)\ to create the assumption \z:B(a)\. + A more natural form of product elimination. +\ lemma subst_prodE: assumes "p: Prod(A,B)" and "a: A" and "\z. z: B(a) \ c(z): C(z)" shows "c(p`a): C(p`a)" -apply (rule assms ProdE)+ -done + by (rule assms ProdE)+ subsection \Tactics for type checking\ ML \ - local fun is_rigid_elem (Const(@{const_name Elem},_) $ a $ _) = not(is_Var (head_of a)) @@ -307,26 +318,22 @@ in (*Try solving a:A or a=b:A by assumption provided a is rigid!*) -fun test_assume_tac ctxt = SUBGOAL(fn (prem,i) => - if is_rigid_elem (Logic.strip_assums_concl prem) - then assume_tac ctxt i else no_tac) +fun test_assume_tac ctxt = SUBGOAL (fn (prem, i) => + if is_rigid_elem (Logic.strip_assums_concl prem) + then assume_tac ctxt i else no_tac) fun ASSUME ctxt tf i = test_assume_tac ctxt i ORELSE tf i -end; - +end \ -(*For simplification: type formation and checking, - but no equalities between terms*) +text \ + For simplification: type formation and checking, + but no equalities between terms. +\ lemmas routine_rls = form_rls formL_rls refl_type element_rls ML \ -local - val equal_rls = @{thms form_rls} @ @{thms element_rls} @ @{thms intrL_rls} @ - @{thms elimL_rls} @ @{thms refl_elem} -in - fun routine_tac rls ctxt prems = ASSUME ctxt (filt_resolve_from_net_tac ctxt 4 (Tactic.build_net (prems @ rls))); @@ -354,9 +361,9 @@ (*Equality proving: solve a=b:A (where a is rigid) by long rules. *) fun equal_tac ctxt thms = REPEAT_FIRST - (ASSUME ctxt (filt_resolve_from_net_tac ctxt 3 (Tactic.build_net (thms @ equal_rls)))) - -end + (ASSUME ctxt + (filt_resolve_from_net_tac ctxt 3 + (Tactic.build_net (thms @ @{thms form_rls element_rls intrL_rls elimL_rls refl_elem})))) \ method_setup form = \Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\ @@ -367,25 +374,25 @@ subsection \Simplification\ -(*To simplify the type in a goal*) +text \To simplify the type in a goal.\ lemma replace_type: "\B = A; a : A\ \ a : B" -apply (rule equal_types) -apply (rule_tac [2] sym_type) -apply assumption+ -done + apply (rule equal_types) + apply (rule_tac [2] sym_type) + apply assumption+ + done -(*Simplify the parameter of a unary type operator.*) +text \Simplify the parameter of a unary type operator.\ lemma subst_eqtyparg: assumes 1: "a=c : A" and 2: "\z. z:A \ B(z) type" - shows "B(a)=B(c)" -apply (rule subst_typeL) -apply (rule_tac [2] refl_type) -apply (rule 1) -apply (erule 2) -done + shows "B(a) = B(c)" + apply (rule subst_typeL) + apply (rule_tac [2] refl_type) + apply (rule 1) + apply (erule 2) + done -(*Simplification rules for Constructive Type Theory*) +text \Simplification rules for Constructive Type Theory.\ lemmas reduction_rls = comp_rls [THEN trans_elem] ML \ @@ -462,12 +469,12 @@ subsection \The elimination rules for fst/snd\ lemma SumE_fst: "p : Sum(A,B) \ fst(p) : A" -apply (unfold basic_defs) -apply (erule SumE) -apply assumption -done + apply (unfold basic_defs) + apply (erule SumE) + apply assumption + done -(*The first premise must be p:Sum(A,B) !!*) +text \The first premise must be \p:Sum(A,B)\!!.\ lemma SumE_snd: assumes major: "p: Sum(A,B)" and "A type" @@ -476,7 +483,7 @@ apply (unfold basic_defs) apply (rule major [THEN SumE]) apply (rule SumC [THEN subst_eqtyparg, THEN replace_type]) - apply (typechk assms) + apply (typechk assms) done end diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/Main.thy --- a/src/CTT/Main.thy Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/Main.thy Fri Jul 15 15:19:04 2016 +0200 @@ -1,6 +1,6 @@ section \Main includes everything\ theory Main -imports CTT Arith Bool + imports CTT Arith Bool begin end diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/ROOT --- a/src/CTT/ROOT Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/ROOT Fri Jul 15 15:19:04 2016 +0200 @@ -16,12 +16,12 @@ Simon Thompson, Type Theory and Functional Programming (Addison-Wesley, 1991) *} - options [document = false] + options [thy_output_source] theories Main - - (* Examples for Constructive Type Theory *) "ex/Typechecking" "ex/Elimination" "ex/Equality" "ex/Synthesis" + document_files + "root.tex" diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/CTT/document/root.tex Fri Jul 15 15:19:04 2016 +0200 @@ -0,0 +1,21 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[utf8]{inputenc} +\usepackage{isabelle,isabellesym} +\usepackage{pdfsetup} + +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Isabelle/CTT --- Constructive Type Theory \\ + with extensional equality and without universes} +\author{Larry Paulson} +\maketitle + +\tableofcontents + +\parindent 0pt\parskip 0.5ex +\input{session} + +\end{document} diff -r 7f0e36eb73b4 -r 42e1dece537a src/CTT/ex/Elimination.thy --- a/src/CTT/ex/Elimination.thy Fri Jul 15 11:26:40 2016 +0200 +++ b/src/CTT/ex/Elimination.thy Fri Jul 15 15:19:04 2016 +0200 @@ -173,7 +173,7 @@ apply assumption done -text "Example of sequent_style deduction" +text "Example of sequent-style deduction" (*When splitting z:A \ B, the assumption C(z) is affected; ?a becomes \<^bold>\u. split(u,\v w.split(v,\x y.\<^bold> \z. >) ` w) *) schematic_goal