--- a/src/CTT/Arith.thy Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/Arith.thy Fri Jul 15 15:19:12 2016 +0200
@@ -6,141 +6,115 @@
section \<open>Elementary arithmetic\<close>
theory Arith
-imports Bool
+ imports Bool
begin
subsection \<open>Arithmetic operators and their definitions\<close>
-definition
- add :: "[i,i]\<Rightarrow>i" (infixr "#+" 65) where
- "a#+b \<equiv> rec(a, b, \<lambda>u v. succ(v))"
+definition add :: "[i,i]\<Rightarrow>i" (infixr "#+" 65)
+ where "a#+b \<equiv> rec(a, b, \<lambda>u v. succ(v))"
-definition
- diff :: "[i,i]\<Rightarrow>i" (infixr "-" 65) where
- "a-b \<equiv> rec(b, a, \<lambda>u v. rec(v, 0, \<lambda>x y. x))"
+definition diff :: "[i,i]\<Rightarrow>i" (infixr "-" 65)
+ where "a-b \<equiv> rec(b, a, \<lambda>u v. rec(v, 0, \<lambda>x y. x))"
-definition
- absdiff :: "[i,i]\<Rightarrow>i" (infixr "|-|" 65) where
- "a|-|b \<equiv> (a-b) #+ (b-a)"
+definition absdiff :: "[i,i]\<Rightarrow>i" (infixr "|-|" 65)
+ where "a|-|b \<equiv> (a-b) #+ (b-a)"
-definition
- mult :: "[i,i]\<Rightarrow>i" (infixr "#*" 70) where
- "a#*b \<equiv> rec(a, 0, \<lambda>u v. b #+ v)"
+definition mult :: "[i,i]\<Rightarrow>i" (infixr "#*" 70)
+ where "a#*b \<equiv> rec(a, 0, \<lambda>u v. b #+ v)"
-definition
- mod :: "[i,i]\<Rightarrow>i" (infixr "mod" 70) where
- "a mod b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(v) |-| b, 0, \<lambda>x y. succ(v)))"
+definition mod :: "[i,i]\<Rightarrow>i" (infixr "mod" 70)
+ where "a mod b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(v) |-| b, 0, \<lambda>x y. succ(v)))"
-definition
- div :: "[i,i]\<Rightarrow>i" (infixr "div" 70) where
- "a div b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(u) mod b, succ(v), \<lambda>x y. v))"
-
+definition div :: "[i,i]\<Rightarrow>i" (infixr "div" 70)
+ where "a div b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(u) mod b, succ(v), \<lambda>x y. v))"
lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def
subsection \<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
-(** Addition *)
+subsubsection \<open>Addition\<close>
-(*typing of add: short and long versions*)
+text \<open>Typing of \<open>add\<close>: short and long versions.\<close>
lemma add_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #+ b : N"
-apply (unfold arith_defs)
-apply typechk
-done
+ unfolding arith_defs by typechk
lemma add_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> 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 \<open>Computation for \<open>add\<close>: 0 and successor cases.\<close>
lemma addC0: "b:N \<Longrightarrow> 0 #+ b = b : N"
-apply (unfold arith_defs)
-apply rew
-done
+ unfolding arith_defs by rew
lemma addC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #+ b = succ(a #+ b) : N"
-apply (unfold arith_defs)
-apply rew
-done
+ unfolding arith_defs by rew
-(** Multiplication *)
+subsubsection \<open>Multiplication\<close>
-(*typing of mult: short and long versions*)
+text \<open>Typing of \<open>mult\<close>: short and long versions.\<close>
lemma mult_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* b : N"
-apply (unfold arith_defs)
-apply (typechk add_typing)
-done
+ unfolding arith_defs by (typechk add_typing)
lemma mult_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> 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 \<open>Computation for \<open>mult\<close>: 0 and successor cases.\<close>
lemma multC0: "b:N \<Longrightarrow> 0 #* b = 0 : N"
-apply (unfold arith_defs)
-apply rew
-done
+ unfolding arith_defs by rew
lemma multC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #* b = b #+ (a #* b) : N"
-apply (unfold arith_defs)
-apply rew
-done
+ unfolding arith_defs by rew
-(** Difference *)
+subsubsection \<open>Difference\<close>
-(*typing of difference*)
+text \<open>Typing of difference.\<close>
lemma diff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a - b : N"
-apply (unfold arith_defs)
-apply typechk
-done
+ unfolding arith_defs by typechk
lemma diff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> 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 \<open>Computation for difference: 0 and successor cases.\<close>
lemma diffC0: "a:N \<Longrightarrow> a - 0 = a : N"
-apply (unfold arith_defs)
-apply rew
-done
+ unfolding arith_defs by rew
-(*Note: rec(a, 0, \<lambda>z w.z) is pred(a). *)
+text \<open>Note: \<open>rec(a, 0, \<lambda>z w.z)\<close> is \<open>pred(a).\<close>\<close>
lemma diff_0_eq_0: "b:N \<Longrightarrow> 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 \<open>
+ Essential to simplify FIRST!! (Else we get a critical pair)
+ \<open>succ(a) - succ(b)\<close> rewrites to \<open>pred(succ(a) - b)\<close>.
+\<close>
lemma diff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Simplification\<close>
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 \<open>
+ 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))
\<close>
method_setup arith_rew = \<open>
@@ -186,284 +153,263 @@
subsection \<open>Addition\<close>
-(*Associative law for addition*)
+text \<open>Associative law for addition.\<close>
lemma add_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (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 \<open>Commutative law for addition. Can be proved using three inductions.
+ Must simplify after first induction! Orientation of rewrites is delicate.\<close>
lemma add_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Multiplication\<close>
-(*right annihilation in product*)
+text \<open>Right annihilation in product.\<close>
lemma mult_0_right: "a:N \<Longrightarrow> 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 \<open>Right successor law for multiplication.\<close>
lemma mult_succ_right: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Commutative law for multiplication.\<close>
lemma mult_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Addition distributes over multiplication.\<close>
lemma add_mult_distrib: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (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 \<open>Associative law for multiplication.\<close>
lemma mult_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (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 \<open>Difference\<close>
text \<open>
-Difference on natural numbers, without negative numbers
- a - b = 0 iff a<=b a - b = succ(c) iff a>b\<close>
+ Difference on natural numbers, without negative numbers
+ \<^item> \<open>a - b = 0\<close> iff \<open>a \<le> b\<close>
+ \<^item> \<open>a - b = succ(c)\<close> iff \<open>a > b\<close>
+\<close>
lemma diff_self_eq_0: "a:N \<Longrightarrow> a - a = 0 : N"
-apply (NE a)
-apply hyp_arith_rew
-done
+ apply (NE a)
+ apply hyp_arith_rew
+ done
lemma add_0_right: "\<lbrakk>c : N; 0 : N; c : N\<rbrakk> \<Longrightarrow> 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 \<open>
+ Addition is the inverse of subtraction: if \<open>b \<le> x\<close> then \<open>b #+ (x - b) = x\<close>.
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.
+\<close>
schematic_goal add_diff_inverse_lemma:
"b:N \<Longrightarrow> ?a : \<Prod>x:N. Eq(N, b-x, 0) \<longrightarrow> 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) \<longrightarrow> (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)
+ \<comment> \<open>strip one "universal quantifier" but not the "implication"\<close>
+ apply (rule_tac [3] intr_rls)
+ \<comment> \<open>case analysis on \<open>x\<close> in \<open>succ(u) \<le> x \<longrightarrow> succ(u) #+ (x - succ(u)) = x\<close>\<close>
+ prefer 4
+ apply (NE x)
+ apply assumption
+ \<comment> \<open>Prepare for simplification of types -- the antecedent \<open>succ(u) \<le> x\<close>\<close>
+ apply (rule_tac [2] replace_type)
+ apply (rule_tac [1] replace_type)
+ apply arith_rew
+ \<comment> \<open>Solves first 0 goal, simplifies others. Two sugbgoals remain.
+ Both follow by rewriting, (2) using quantified induction hyp.\<close>
+ apply intr \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
+ 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 \<open>
+ Version of above with premise \<open>b - a = 0\<close> i.e. \<open>a \<ge> b\<close>.
+ Using @{thm ProdE} does not work -- for \<open>?B(?a)\<close> is ambiguous.
+ Instead, @{thm add_diff_inverse_lemma} states the desired induction scheme;
+ the use of \<open>THEN\<close> below instantiates Vars in @{thm ProdE} automatically.
+\<close>
lemma add_diff_inverse: "\<lbrakk>a:N; b:N; b - a = 0 : N\<rbrakk> \<Longrightarrow> 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 \<open>Absolute difference\<close>
-(*typing of absolute difference: short and long versions*)
+text \<open>Typing of absolute difference: short and long versions.\<close>
lemma absdiff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a |-| b : N"
-apply (unfold arith_defs)
-apply typechk
-done
+ unfolding arith_defs by typechk
lemma absdiff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a |-| b = c |-| d : N"
-apply (unfold arith_defs)
-apply equal
-done
+ unfolding arith_defs by equal
lemma absdiff_self_eq_0: "a:N \<Longrightarrow> 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 \<Longrightarrow> 0 |-| a = a : N"
-apply (unfold absdiff_def)
-apply hyp_arith_rew
-done
-
+ unfolding absdiff_def by hyp_arith_rew
lemma absdiff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Note how easy using commutative laws can be? ...not always...\<close>
lemma absdiff_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>If \<open>a + b = 0\<close> then \<open>a = 0\<close>. Surprisingly tedious.\<close>
schematic_goal add_eq0_lemma: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> ?c : \<Prod>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 \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
+ 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 \<open>
+ Version of above with the premise \<open>a + b = 0\<close>.
+ Again, resolution instantiates variables in @{thm ProdE}.
+\<close>
lemma add_eq0: "\<lbrakk>a:N; b:N; a #+ b = 0 : N\<rbrakk> \<Longrightarrow> 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 \<open>Here is a lemma to infer \<open>a - b = 0\<close> and \<open>b - a = 0\<close> from \<open>a |-| b = 0\<close>, below.\<close>
schematic_goal absdiff_eq0_lem:
"\<lbrakk>a:N; b:N; a |-| b = 0 : N\<rbrakk> \<Longrightarrow> ?a : \<Sum>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 \<open>If \<open>a |-| b = 0\<close> then \<open>a = b\<close>
+ proof: \<open>a - b = 0\<close> and \<open>b - a = 0\<close>, so \<open>b = a + (b - a) = a + 0 = a\<close>.
+\<close>
lemma absdiff_eq0: "\<lbrakk>a |-| b = 0 : N; a:N; b:N\<rbrakk> \<Longrightarrow> 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 \<open>Remainder and Quotient\<close>
-(*typing of remainder: short and long versions*)
+text \<open>Typing of remainder: short and long versions.\<close>
lemma mod_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a mod b : N"
-apply (unfold mod_def)
-apply (typechk absdiff_typing)
-done
+ unfolding mod_def by (typechk absdiff_typing)
lemma mod_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> 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 \<open>Computation for \<open>mod\<close>: 0 and successor cases.\<close>
lemma modC0: "b:N \<Longrightarrow> 0 mod b = 0 : N"
-apply (unfold mod_def)
-apply (rew absdiff_typing)
-done
+ unfolding mod_def by (rew absdiff_typing)
lemma modC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
succ(a) mod b = rec(succ(a mod b) |-| b, 0, \<lambda>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 \<open>Typing of quotient: short and long versions.\<close>
lemma div_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> 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 \<open>Computation for quotient: 0 and successor cases.\<close>
lemma divC0: "b:N \<Longrightarrow> 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: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
succ(a) div b = rec(succ(a) mod b, succ(a div b), \<lambda>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 \<open>Version of above with same condition as the \<open>mod\<close> one.\<close>
lemma divC_succ2: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), \<lambda>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 \<open>For case analysis on whether a number is 0 or a successor.\<close>
lemma iszero_decidable: "a:N \<Longrightarrow> rec(a, inl(eq), \<lambda>ka kb. inr(<ka, eq>)) :
- Eq(N,a,0) + (\<Sum>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) + (\<Sum>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 \<open>Main Result. Holds when \<open>b\<close> is 0 since \<open>a mod 0 = a\<close> and \<open>a div 0 = 0\<close>.\<close>
lemma mod_div_equality: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> 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)
+ \<comment> \<open>case analysis on \<open>succ(u mod b) |-| b\<close>\<close>
+ 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)
+ \<comment> \<open>Replace one occurrence of \<open>b\<close> by \<open>succ(u mod b)\<close>. Clumsy!\<close>
+ 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
--- a/src/CTT/Bool.thy Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/Bool.thy Fri Jul 15 15:19:12 2016 +0200
@@ -6,80 +6,68 @@
section \<open>The two-element type (booleans and conditionals)\<close>
theory Bool
-imports CTT
+ imports CTT
begin
-definition
- Bool :: "t" where
- "Bool \<equiv> T+T"
+definition Bool :: "t"
+ where "Bool \<equiv> T+T"
-definition
- true :: "i" where
- "true \<equiv> inl(tt)"
+definition true :: "i"
+ where "true \<equiv> inl(tt)"
-definition
- false :: "i" where
- "false \<equiv> inr(tt)"
+definition false :: "i"
+ where "false \<equiv> inr(tt)"
-definition
- cond :: "[i,i,i]\<Rightarrow>i" where
- "cond(a,b,c) \<equiv> when(a, \<lambda>u. b, \<lambda>u. c)"
+definition cond :: "[i,i,i]\<Rightarrow>i"
+ where "cond(a,b,c) \<equiv> when(a, \<lambda>u. b, \<lambda>u. c)"
lemmas bool_defs = Bool_def true_def false_def cond_def
-subsection \<open>Derivation of rules for the type Bool\<close>
+subsection \<open>Derivation of rules for the type \<open>Bool\<close>\<close>
-(*formation rule*)
+text \<open>Formation rule.\<close>
lemma boolF: "Bool type"
-apply (unfold bool_defs)
-apply typechk
-done
+ unfolding bool_defs by typechk
-
-(*introduction rules for true, false*)
+text \<open>Introduction rules for \<open>true\<close>, \<open>false\<close>.\<close>
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 \<open>Elimination rule: typing of \<open>cond\<close>.\<close>
lemma boolE: "\<lbrakk>p:Bool; a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>p = q : Bool; a = c : C(true); b = d : C(false)\<rbrakk>
\<Longrightarrow> 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 \<open>Computation rules for \<open>true\<close>, \<open>false\<close>.\<close>
lemma boolC_true: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> 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
--- a/src/CTT/CTT.thy Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/CTT.thy Fri Jul 15 15:19:12 2016 +0200
@@ -17,45 +17,46 @@
typedecl o
consts
- (*Types*)
+ \<comment> \<open>Types\<close>
F :: "t"
- T :: "t" (*F is empty, T contains one element*)
+ T :: "t" \<comment> \<open>\<open>F\<close> is empty, \<open>T\<close> contains one element\<close>
contr :: "i\<Rightarrow>i"
tt :: "i"
- (*Natural numbers*)
+ \<comment> \<open>Natural numbers\<close>
N :: "t"
succ :: "i\<Rightarrow>i"
rec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow> i"
- (*Unions*)
+ \<comment> \<open>Unions\<close>
inl :: "i\<Rightarrow>i"
inr :: "i\<Rightarrow>i"
"when" :: "[i, i\<Rightarrow>i, i\<Rightarrow>i]\<Rightarrow>i"
- (*General Sum and Binary Product*)
+ \<comment> \<open>General Sum and Binary Product\<close>
Sum :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
fst :: "i\<Rightarrow>i"
snd :: "i\<Rightarrow>i"
split :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"
- (*General Product and Function Space*)
+ \<comment> \<open>General Product and Function Space\<close>
Prod :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
- (*Types*)
+ \<comment> \<open>Types\<close>
Plus :: "[t,t]\<Rightarrow>t" (infixr "+" 40)
- (*Equality type*)
+ \<comment> \<open>Equality type\<close>
Eq :: "[t,i,i]\<Rightarrow>t"
eq :: "i"
- (*Judgements*)
+ \<comment> \<open>Judgements\<close>
Type :: "t \<Rightarrow> prop" ("(_ type)" [10] 5)
Eqtype :: "[t,t]\<Rightarrow>prop" ("(_ =/ _)" [10,10] 5)
Elem :: "[i, t]\<Rightarrow>prop" ("(_ /: _)" [10,10] 5)
Eqelem :: "[i,i,t]\<Rightarrow>prop" ("(_ =/ _ :/ _)" [10,10,10] 5)
Reduce :: "[i,i]\<Rightarrow>prop" ("Reduce[_,_]")
- (*Types*)
- (*Functions*)
+ \<comment> \<open>Types\<close>
+
+ \<comment> \<open>Functions\<close>
lambda :: "(i \<Rightarrow> i) \<Rightarrow> i" (binder "\<^bold>\<lambda>" 10)
app :: "[i,i]\<Rightarrow>i" (infixl "`" 60)
- (*Natural numbers*)
+ \<comment> \<open>Natural numbers\<close>
Zero :: "i" ("0")
- (*Pairing*)
+ \<comment> \<open>Pairing\<close>
pair :: "[i,i]\<Rightarrow>i" ("(1<_,/_>)")
syntax
@@ -65,35 +66,37 @@
"\<Prod>x:A. B" \<rightleftharpoons> "CONST Prod(A, \<lambda>x. B)"
"\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum(A, \<lambda>x. B)"
-abbreviation
- Arrow :: "[t,t]\<Rightarrow>t" (infixr "\<longrightarrow>" 30) where
- "A \<longrightarrow> B \<equiv> \<Prod>_:A. B"
-abbreviation
- Times :: "[t,t]\<Rightarrow>t" (infixr "\<times>" 50) where
- "A \<times> B \<equiv> \<Sum>_:A. B"
+abbreviation Arrow :: "[t,t]\<Rightarrow>t" (infixr "\<longrightarrow>" 30)
+ where "A \<longrightarrow> B \<equiv> \<Prod>_:A. B"
+
+abbreviation Times :: "[t,t]\<Rightarrow>t" (infixr "\<times>" 50)
+ where "A \<times> B \<equiv> \<Sum>_: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 \<open>
+ Reduction: a weaker notion than equality; a hack for simplification.
+ \<open>Reduce[a,b]\<close> means either that \<open>a = b : A\<close> for some \<open>A\<close> or else
+ that \<open>a\<close> and \<open>b\<close> 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 \<open>a:A\<close>! Sound because only \<open>trans_red\<close> uses a \<open>Reduce\<close>
+ premise. No new theorems can be proved about the standard judgements.
+\<close>
+axiomatization
+where
refl_red: "\<And>a. Reduce[a,a]" and
red_if_equal: "\<And>a b A. a = b : A \<Longrightarrow> Reduce[a,b]" and
trans_red: "\<And>a b c A. \<lbrakk>a = b : A; Reduce[b,c]\<rbrakk> \<Longrightarrow> a = c : A" and
- (*Reflexivity*)
+ \<comment> \<open>Reflexivity\<close>
refl_type: "\<And>A. A type \<Longrightarrow> A = A" and
refl_elem: "\<And>a A. a : A \<Longrightarrow> a = a : A" and
- (*Symmetry*)
+ \<comment> \<open>Symmetry\<close>
sym_type: "\<And>A B. A = B \<Longrightarrow> B = A" and
sym_elem: "\<And>a b A. a = b : A \<Longrightarrow> b = a : A" and
- (*Transitivity*)
+ \<comment> \<open>Transitivity\<close>
trans_type: "\<And>A B C. \<lbrakk>A = B; B = C\<rbrakk> \<Longrightarrow> A = C" and
trans_elem: "\<And>a b c A. \<lbrakk>a = b : A; b = c : A\<rbrakk> \<Longrightarrow> a = c : A" and
@@ -101,7 +104,7 @@
equal_types: "\<And>a A B. \<lbrakk>a : A; A = B\<rbrakk> \<Longrightarrow> a : B" and
equal_typesL: "\<And>a b A B. \<lbrakk>a = b : A; A = B\<rbrakk> \<Longrightarrow> a = b : B" and
- (*Substitution*)
+ \<comment> \<open>Substitution\<close>
subst_type: "\<And>a A B. \<lbrakk>a : A; \<And>z. z:A \<Longrightarrow> B(z) type\<rbrakk> \<Longrightarrow> B(a) type" and
subst_typeL: "\<And>a c A B D. \<lbrakk>a = c : A; \<And>z. z:A \<Longrightarrow> B(z) = D(z)\<rbrakk> \<Longrightarrow> B(a) = D(c)" and
@@ -111,7 +114,7 @@
"\<And>a b c d A B. \<lbrakk>a = c : A; \<And>z. z:A \<Longrightarrow> b(z)=d(z) : B(z)\<rbrakk> \<Longrightarrow> b(a)=d(c) : B(a)" and
- (*The type N -- natural numbers*)
+ \<comment> \<open>The type \<open>N\<close> -- natural numbers\<close>
NF: "N type" and
NI0: "0 : N" and
@@ -135,11 +138,11 @@
"\<And>p a b C. \<lbrakk>p: N; a: C(0); \<And>u v. \<lbrakk>u: N; v: C(u)\<rbrakk> \<Longrightarrow> b(u,v): C(succ(u))\<rbrakk> \<Longrightarrow>
rec(succ(p), a, \<lambda>u v. b(u,v)) = b(p, rec(p, a, \<lambda>u v. b(u,v))) : C(succ(p))" and
- (*The fourth Peano axiom. See page 91 of Martin-Löf's book*)
+ \<comment> \<open>The fourth Peano axiom. See page 91 of Martin-Löf's book.\<close>
zero_ne_succ: "\<And>a. \<lbrakk>a: N; 0 = succ(a) : N\<rbrakk> \<Longrightarrow> 0: F" and
- (*The Product of a family of types*)
+ \<comment> \<open>The Product of a family of types\<close>
ProdF: "\<And>A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> B(x) type\<rbrakk> \<Longrightarrow> \<Prod>x:A. B(x) type" and
@@ -160,7 +163,7 @@
ProdC2: "\<And>p A B. p : \<Prod>x:A. B(x) \<Longrightarrow> (\<^bold>\<lambda>x. p`x) = p : \<Prod>x:A. B(x)" and
- (*The Sum of a family of types*)
+ \<comment> \<open>The Sum of a family of types\<close>
SumF: "\<And>A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> B(x) type\<rbrakk> \<Longrightarrow> \<Sum>x:A. B(x) type" and
SumFL: "\<And>A B C D. \<lbrakk>A = C; \<And>x. x:A \<Longrightarrow> B(x) = D(x)\<rbrakk> \<Longrightarrow> \<Sum>x:A. B(x) = \<Sum>x:C. D(x)" and
@@ -182,7 +185,7 @@
snd_def: "\<And>a. snd(a) \<equiv> split(a, \<lambda>x y. y)" and
- (*The sum of two types*)
+ \<comment> \<open>The sum of two types\<close>
PlusF: "\<And>A B. \<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> A+B type" and
PlusFL: "\<And>A B C D. \<lbrakk>A = C; B = D\<rbrakk> \<Longrightarrow> A+B = C+D" and
@@ -217,27 +220,31 @@
\<Longrightarrow> when(inr(b), \<lambda>x. c(x), \<lambda>y. d(y)) = d(b) : C(inr(b))" and
- (*The type Eq*)
+ \<comment> \<open>The type \<open>Eq\<close>\<close>
EqF: "\<And>a b A. \<lbrakk>A type; a : A; b : A\<rbrakk> \<Longrightarrow> Eq(A,a,b) type" and
EqFL: "\<And>a b c d A B. \<lbrakk>A = B; a = c : A; b = d : A\<rbrakk> \<Longrightarrow> Eq(A,a,b) = Eq(B,c,d)" and
EqI: "\<And>a b A. a = b : A \<Longrightarrow> eq : Eq(A,a,b)" and
EqE: "\<And>p a b A. p : Eq(A,a,b) \<Longrightarrow> a = b : A" and
- (*By equality of types, can prove C(p) from C(eq), an elimination rule*)
+ \<comment> \<open>By equality of types, can prove \<open>C(p)\<close> from \<open>C(eq)\<close>, an elimination rule\<close>
EqC: "\<And>p a b A. p : Eq(A,a,b) \<Longrightarrow> p = eq : Eq(A,a,b)" and
- (*The type F*)
+
+ \<comment> \<open>The type \<open>F\<close>\<close>
FF: "F type" and
FE: "\<And>p C. \<lbrakk>p: F; C type\<rbrakk> \<Longrightarrow> contr(p) : C" and
FEL: "\<And>p q C. \<lbrakk>p = q : F; C type\<rbrakk> \<Longrightarrow> 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. *)
+
+ \<comment> \<open>The type T\<close>
+ \<comment> \<open>
+ 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 \<open>C(x)\<close> type \<open>x:T\<close>.
+ Also computation can be derived from elimination.
+ \<close>
TF: "T type" and
TI: "tt : T" and
@@ -248,55 +255,59 @@
subsection "Tactics and derived rules for Constructive Type Theory"
-(*Formation rules*)
+text \<open>Formation rules.\<close>
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 \<open>
+ Introduction rules. OMITTED:
+ \<^item> \<open>EqI\<close>, because its premise is an \<open>eqelem\<close>, not an \<open>elem\<close>.
+\<close>
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 \<open>
+ Elimination rules. OMITTED:
+ \<^item> \<open>EqE\<close>, because its conclusion is an \<open>eqelem\<close>, not an \<open>elem\<close>
+ \<^item> \<open>TE\<close>, because it does not involve a constructor.
+\<close>
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 \<open>OMITTED: \<open>eqC\<close> are \<open>TC\<close> because they make rewriting loop: \<open>p = un = un = \<dots>\<close>\<close>
lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr
-(*rules with conclusion a:A, an elem judgement*)
+text \<open>Rules with conclusion \<open>a:A\<close>, an elem judgement.\<close>
lemmas element_rls = intr_rls elim_rls
-(*Definitions are (meta)equality axioms*)
+text \<open>Definitions are (meta)equality axioms.\<close>
lemmas basic_defs = fst_def snd_def
-(*Compare with standard version: B is applied to UNSIMPLIFIED expression! *)
+text \<open>Compare with standard version: \<open>B\<close> is applied to UNSIMPLIFIED expression!\<close>
lemma SumIL2: "\<lbrakk>c = a : A; d = b : B(a)\<rbrakk> \<Longrightarrow> <c,d> = <a,b> : 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 \<open>
+ Exploit \<open>p:Prod(A,B)\<close> to create the assumption \<open>z:B(a)\<close>.
+ A more natural form of product elimination.
+\<close>
lemma subst_prodE:
assumes "p: Prod(A,B)"
and "a: A"
and "\<And>z. z: B(a) \<Longrightarrow> c(z): C(z)"
shows "c(p`a): C(p`a)"
-apply (rule assms ProdE)+
-done
+ by (rule assms ProdE)+
subsection \<open>Tactics for type checking\<close>
ML \<open>
-
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
\<close>
-(*For simplification: type formation and checking,
- but no equalities between terms*)
+text \<open>
+ For simplification: type formation and checking,
+ but no equalities between terms.
+\<close>
lemmas routine_rls = form_rls formL_rls refl_type element_rls
ML \<open>
-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}))))
\<close>
method_setup form = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\<close>
@@ -367,25 +374,25 @@
subsection \<open>Simplification\<close>
-(*To simplify the type in a goal*)
+text \<open>To simplify the type in a goal.\<close>
lemma replace_type: "\<lbrakk>B = A; a : A\<rbrakk> \<Longrightarrow> 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 \<open>Simplify the parameter of a unary type operator.\<close>
lemma subst_eqtyparg:
assumes 1: "a=c : A"
and 2: "\<And>z. z:A \<Longrightarrow> 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 \<open>Simplification rules for Constructive Type Theory.\<close>
lemmas reduction_rls = comp_rls [THEN trans_elem]
ML \<open>
@@ -462,12 +469,12 @@
subsection \<open>The elimination rules for fst/snd\<close>
lemma SumE_fst: "p : Sum(A,B) \<Longrightarrow> 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 \<open>The first premise must be \<open>p:Sum(A,B)\<close>!!.\<close>
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
--- a/src/CTT/Main.thy Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/Main.thy Fri Jul 15 15:19:12 2016 +0200
@@ -1,6 +1,6 @@
section \<open>Main includes everything\<close>
theory Main
-imports CTT Arith Bool
+ imports CTT Arith Bool
begin
end
--- a/src/CTT/ROOT Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/ROOT Fri Jul 15 15:19:12 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"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/CTT/document/root.tex Fri Jul 15 15:19:12 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}
--- a/src/CTT/ex/Elimination.thy Fri Jul 15 13:45:19 2016 +0200
+++ b/src/CTT/ex/Elimination.thy Fri Jul 15 15:19:12 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 \<times> B, the assumption C(z) is affected; ?a becomes
\<^bold>\<lambda>u. split(u,\<lambda>v w.split(v,\<lambda>x y.\<^bold> \<lambda>z. <x,<y,z>>) ` w) *)
schematic_goal