1(*  Title:      CTT/CTT.thy
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3    Copyright   1993  University of Cambridge
4*)
5
6theory CTT
7imports Pure
8begin
9
10section \<open>Constructive Type Theory: axiomatic basis\<close>
11
12ML_file \<open>~~/src/Provers/typedsimp.ML\<close>
13setup Pure_Thy.old_appl_syntax_setup
14
15typedecl i
16typedecl t
17typedecl o
18
19consts
20  \<comment> \<open>Types\<close>
21  F         :: "t"
22  T         :: "t"          \<comment> \<open>\<open>F\<close> is empty, \<open>T\<close> contains one element\<close>
23  contr     :: "i\<Rightarrow>i"
24  tt        :: "i"
25  \<comment> \<open>Natural numbers\<close>
26  N         :: "t"
27  succ      :: "i\<Rightarrow>i"
28  rec       :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow> i"
29  \<comment> \<open>Unions\<close>
30  inl       :: "i\<Rightarrow>i"
31  inr       :: "i\<Rightarrow>i"
32  "when"    :: "[i, i\<Rightarrow>i, i\<Rightarrow>i]\<Rightarrow>i"
33  \<comment> \<open>General Sum and Binary Product\<close>
34  Sum       :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
35  fst       :: "i\<Rightarrow>i"
36  snd       :: "i\<Rightarrow>i"
37  split     :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"
38  \<comment> \<open>General Product and Function Space\<close>
39  Prod      :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
40  \<comment> \<open>Types\<close>
41  Plus      :: "[t,t]\<Rightarrow>t"           (infixr "+" 40)
42  \<comment> \<open>Equality type\<close>
43  Eq        :: "[t,i,i]\<Rightarrow>t"
44  eq        :: "i"
45  \<comment> \<open>Judgements\<close>
46  Type      :: "t \<Rightarrow> prop"          ("(_ type)" [10] 5)
47  Eqtype    :: "[t,t]\<Rightarrow>prop"        ("(_ =/ _)" [10,10] 5)
48  Elem      :: "[i, t]\<Rightarrow>prop"       ("(_ /: _)" [10,10] 5)
49  Eqelem    :: "[i,i,t]\<Rightarrow>prop"      ("(_ =/ _ :/ _)" [10,10,10] 5)
50  Reduce    :: "[i,i]\<Rightarrow>prop"        ("Reduce[_,_]")
51
52  \<comment> \<open>Types\<close>
53
54  \<comment> \<open>Functions\<close>
55  lambda    :: "(i \<Rightarrow> i) \<Rightarrow> i"      (binder "\<^bold>\<lambda>" 10)
56  app       :: "[i,i]\<Rightarrow>i"           (infixl "`" 60)
57  \<comment> \<open>Natural numbers\<close>
58  Zero      :: "i"                  ("0")
59  \<comment> \<open>Pairing\<close>
60  pair      :: "[i,i]\<Rightarrow>i"           ("(1<_,/_>)")
61
62syntax
63  "_PROD"   :: "[idt,t,t]\<Rightarrow>t"       ("(3\<Prod>_:_./ _)" 10)
64  "_SUM"    :: "[idt,t,t]\<Rightarrow>t"       ("(3\<Sum>_:_./ _)" 10)
65translations
66  "\<Prod>x:A. B" \<rightleftharpoons> "CONST Prod(A, \<lambda>x. B)"
67  "\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum(A, \<lambda>x. B)"
68
69abbreviation Arrow :: "[t,t]\<Rightarrow>t"  (infixr "\<longrightarrow>" 30)
70  where "A \<longrightarrow> B \<equiv> \<Prod>_:A. B"
71
72abbreviation Times :: "[t,t]\<Rightarrow>t"  (infixr "\<times>" 50)
73  where "A \<times> B \<equiv> \<Sum>_:A. B"
74
75text \<open>
76  Reduction: a weaker notion than equality;  a hack for simplification.
77  \<open>Reduce[a,b]\<close> means either that \<open>a = b : A\<close> for some \<open>A\<close> or else
78    that \<open>a\<close> and \<open>b\<close> are textually identical.
79
80  Does not verify \<open>a:A\<close>!  Sound because only \<open>trans_red\<close> uses a \<open>Reduce\<close>
81  premise. No new theorems can be proved about the standard judgements.
82\<close>
83axiomatization
84where
85  refl_red: "\<And>a. Reduce[a,a]" and
86  red_if_equal: "\<And>a b A. a = b : A \<Longrightarrow> Reduce[a,b]" and
87  trans_red: "\<And>a b c A. \<lbrakk>a = b : A; Reduce[b,c]\<rbrakk> \<Longrightarrow> a = c : A" and
88
89  \<comment> \<open>Reflexivity\<close>
90
91  refl_type: "\<And>A. A type \<Longrightarrow> A = A" and
92  refl_elem: "\<And>a A. a : A \<Longrightarrow> a = a : A" and
93
94  \<comment> \<open>Symmetry\<close>
95
96  sym_type:  "\<And>A B. A = B \<Longrightarrow> B = A" and
97  sym_elem:  "\<And>a b A. a = b : A \<Longrightarrow> b = a : A" and
98
99  \<comment> \<open>Transitivity\<close>
100
101  trans_type:   "\<And>A B C. \<lbrakk>A = B; B = C\<rbrakk> \<Longrightarrow> A = C" and
102  trans_elem:   "\<And>a b c A. \<lbrakk>a = b : A; b = c : A\<rbrakk> \<Longrightarrow> a = c : A" and
103
104  equal_types:  "\<And>a A B. \<lbrakk>a : A; A = B\<rbrakk> \<Longrightarrow> a : B" and
105  equal_typesL: "\<And>a b A B. \<lbrakk>a = b : A; A = B\<rbrakk> \<Longrightarrow> a = b : B" and
106
107  \<comment> \<open>Substitution\<close>
108
109  subst_type:   "\<And>a A B. \<lbrakk>a : A; \<And>z. z:A \<Longrightarrow> B(z) type\<rbrakk> \<Longrightarrow> B(a) type" and
110  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
112  subst_elem:   "\<And>a b A B. \<lbrakk>a : A; \<And>z. z:A \<Longrightarrow> b(z):B(z)\<rbrakk> \<Longrightarrow> b(a):B(a)" and
113  subst_elemL:
114    "\<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
115
116
117  \<comment> \<open>The type \<open>N\<close> -- natural numbers\<close>
118
119  NF: "N type" and
120  NI0: "0 : N" and
121  NI_succ: "\<And>a. a : N \<Longrightarrow> succ(a) : N" and
122  NI_succL:  "\<And>a b. a = b : N \<Longrightarrow> succ(a) = succ(b) : N" and
123
124  NE:
125   "\<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>
126   \<Longrightarrow> rec(p, a, \<lambda>u v. b(u,v)) : C(p)" and
127
128  NEL:
129   "\<And>p q a b c d C. \<lbrakk>p = q : N; a = c : C(0);
130      \<And>u v. \<lbrakk>u: N; v: C(u)\<rbrakk> \<Longrightarrow> b(u,v) = d(u,v): C(succ(u))\<rbrakk>
131   \<Longrightarrow> rec(p, a, \<lambda>u v. b(u,v)) = rec(q,c,d) : C(p)" and
132
133  NC0:
134   "\<And>a b C. \<lbrakk>a: C(0); \<And>u v. \<lbrakk>u: N; v: C(u)\<rbrakk> \<Longrightarrow> b(u,v): C(succ(u))\<rbrakk>
135   \<Longrightarrow> rec(0, a, \<lambda>u v. b(u,v)) = a : C(0)" and
136
137  NC_succ:
138   "\<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>
139   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
140
141  \<comment> \<open>The fourth Peano axiom.  See page 91 of Martin-L��f's book.\<close>
142  zero_ne_succ: "\<And>a. \<lbrakk>a: N; 0 = succ(a) : N\<rbrakk> \<Longrightarrow> 0: F" and
143
144
145  \<comment> \<open>The Product of a family of types\<close>
146
147  ProdF: "\<And>A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> B(x) type\<rbrakk> \<Longrightarrow> \<Prod>x:A. B(x) type" and
148
149  ProdFL:
150    "\<And>A B C D. \<lbrakk>A = C; \<And>x. x:A \<Longrightarrow> B(x) = D(x)\<rbrakk> \<Longrightarrow> \<Prod>x:A. B(x) = \<Prod>x:C. D(x)" and
151
152  ProdI:
153    "\<And>b A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> b(x):B(x)\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b(x) : \<Prod>x:A. B(x)" and
154
155  ProdIL: "\<And>b c A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> b(x) = c(x) : B(x)\<rbrakk> \<Longrightarrow>
156    \<^bold>\<lambda>x. b(x) = \<^bold>\<lambda>x. c(x) : \<Prod>x:A. B(x)" and
157
158  ProdE:  "\<And>p a A B. \<lbrakk>p : \<Prod>x:A. B(x); a : A\<rbrakk> \<Longrightarrow> p`a : B(a)" and
159  ProdEL: "\<And>p q a b A B. \<lbrakk>p = q: \<Prod>x:A. B(x); a = b : A\<rbrakk> \<Longrightarrow> p`a = q`b : B(a)" and
160
161  ProdC: "\<And>a b A B. \<lbrakk>a : A; \<And>x. x:A \<Longrightarrow> b(x) : B(x)\<rbrakk> \<Longrightarrow> (\<^bold>\<lambda>x. b(x)) ` a = b(a) : B(a)" and
162
163  ProdC2: "\<And>p A B. p : \<Prod>x:A. B(x) \<Longrightarrow> (\<^bold>\<lambda>x. p`x) = p : \<Prod>x:A. B(x)" and
164
165
166  \<comment> \<open>The Sum of a family of types\<close>
167
168  SumF:  "\<And>A B. \<lbrakk>A type; \<And>x. x:A \<Longrightarrow> B(x) type\<rbrakk> \<Longrightarrow> \<Sum>x:A. B(x) type" and
169  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
170
171  SumI:  "\<And>a b A B. \<lbrakk>a : A; b : B(a)\<rbrakk> \<Longrightarrow> <a,b> : \<Sum>x:A. B(x)" and
172  SumIL: "\<And>a b c d A B. \<lbrakk> a = c : A; b = d : B(a)\<rbrakk> \<Longrightarrow> <a,b> = <c,d> : \<Sum>x:A. B(x)" and
173
174  SumE: "\<And>p c A B C. \<lbrakk>p: \<Sum>x:A. B(x); \<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> c(x,y): C(<x,y>)\<rbrakk>
175    \<Longrightarrow> split(p, \<lambda>x y. c(x,y)) : C(p)" and
176
177  SumEL: "\<And>p q c d A B C. \<lbrakk>p = q : \<Sum>x:A. B(x);
178      \<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> c(x,y)=d(x,y): C(<x,y>)\<rbrakk>
179    \<Longrightarrow> split(p, \<lambda>x y. c(x,y)) = split(q, \<lambda>x y. d(x,y)) : C(p)" and
180
181  SumC: "\<And>a b c A B C. \<lbrakk>a: A;  b: B(a); \<And>x y. \<lbrakk>x:A; y:B(x)\<rbrakk> \<Longrightarrow> c(x,y): C(<x,y>)\<rbrakk>
182    \<Longrightarrow> split(<a,b>, \<lambda>x y. c(x,y)) = c(a,b) : C(<a,b>)" and
183
184  fst_def:   "\<And>a. fst(a) \<equiv> split(a, \<lambda>x y. x)" and
185  snd_def:   "\<And>a. snd(a) \<equiv> split(a, \<lambda>x y. y)" and
186
187
188  \<comment> \<open>The sum of two types\<close>
189
190  PlusF: "\<And>A B. \<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> A+B type" and
191  PlusFL: "\<And>A B C D. \<lbrakk>A = C; B = D\<rbrakk> \<Longrightarrow> A+B = C+D" and
192
193  PlusI_inl: "\<And>a A B. \<lbrakk>a : A; B type\<rbrakk> \<Longrightarrow> inl(a) : A+B" and
194  PlusI_inlL: "\<And>a c A B. \<lbrakk>a = c : A; B type\<rbrakk> \<Longrightarrow> inl(a) = inl(c) : A+B" and
195
196  PlusI_inr: "\<And>b A B. \<lbrakk>A type; b : B\<rbrakk> \<Longrightarrow> inr(b) : A+B" and
197  PlusI_inrL: "\<And>b d A B. \<lbrakk>A type; b = d : B\<rbrakk> \<Longrightarrow> inr(b) = inr(d) : A+B" and
198
199  PlusE:
200    "\<And>p c d A B C. \<lbrakk>p: A+B;
201      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
202      \<And>y. y:B \<Longrightarrow> d(y): C(inr(y)) \<rbrakk> \<Longrightarrow> when(p, \<lambda>x. c(x), \<lambda>y. d(y)) : C(p)" and
203
204  PlusEL:
205    "\<And>p q c d e f A B C. \<lbrakk>p = q : A+B;
206      \<And>x. x: A \<Longrightarrow> c(x) = e(x) : C(inl(x));
207      \<And>y. y: B \<Longrightarrow> d(y) = f(y) : C(inr(y))\<rbrakk>
208    \<Longrightarrow> when(p, \<lambda>x. c(x), \<lambda>y. d(y)) = when(q, \<lambda>x. e(x), \<lambda>y. f(y)) : C(p)" and
209
210  PlusC_inl:
211    "\<And>a c d A B C. \<lbrakk>a: A;
212      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
213      \<And>y. y:B \<Longrightarrow> d(y): C(inr(y)) \<rbrakk>
214    \<Longrightarrow> when(inl(a), \<lambda>x. c(x), \<lambda>y. d(y)) = c(a) : C(inl(a))" and
215
216  PlusC_inr:
217    "\<And>b c d A B C. \<lbrakk>b: B;
218      \<And>x. x:A \<Longrightarrow> c(x): C(inl(x));
219      \<And>y. y:B \<Longrightarrow> d(y): C(inr(y))\<rbrakk>
220    \<Longrightarrow> when(inr(b), \<lambda>x. c(x), \<lambda>y. d(y)) = d(b) : C(inr(b))" and
221
222
223  \<comment> \<open>The type \<open>Eq\<close>\<close>
224
225  EqF: "\<And>a b A. \<lbrakk>A type; a : A; b : A\<rbrakk> \<Longrightarrow> Eq(A,a,b) type" and
226  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
227  EqI: "\<And>a b A. a = b : A \<Longrightarrow> eq : Eq(A,a,b)" and
228  EqE: "\<And>p a b A. p : Eq(A,a,b) \<Longrightarrow> a = b : A" and
229
230  \<comment> \<open>By equality of types, can prove \<open>C(p)\<close> from \<open>C(eq)\<close>, an elimination rule\<close>
231  EqC: "\<And>p a b A. p : Eq(A,a,b) \<Longrightarrow> p = eq : Eq(A,a,b)" and
232
233
234  \<comment> \<open>The type \<open>F\<close>\<close>
235
236  FF: "F type" and
237  FE: "\<And>p C. \<lbrakk>p: F; C type\<rbrakk> \<Longrightarrow> contr(p) : C" and
238  FEL: "\<And>p q C. \<lbrakk>p = q : F; C type\<rbrakk> \<Longrightarrow> contr(p) = contr(q) : C" and
239
240
241  \<comment> \<open>The type T\<close>
242  \<comment> \<open>Martin-L��f's book (page 68) discusses elimination and computation.
243    Elimination can be derived by computation and equality of types,
244    but with an extra premise \<open>C(x)\<close> type \<open>x:T\<close>.
245    Also computation can be derived from elimination.\<close>
246
247  TF: "T type" and
248  TI: "tt : T" and
249  TE: "\<And>p c C. \<lbrakk>p : T; c : C(tt)\<rbrakk> \<Longrightarrow> c : C(p)" and
250  TEL: "\<And>p q c d C. \<lbrakk>p = q : T; c = d : C(tt)\<rbrakk> \<Longrightarrow> c = d : C(p)" and
251  TC: "\<And>p. p : T \<Longrightarrow> p = tt : T"
252
253
254subsection "Tactics and derived rules for Constructive Type Theory"
255
256text \<open>Formation rules.\<close>
257lemmas form_rls = NF ProdF SumF PlusF EqF FF TF
258  and formL_rls = ProdFL SumFL PlusFL EqFL
259
260text \<open>
261  Introduction rules. OMITTED:
262  \<^item> \<open>EqI\<close>, because its premise is an \<open>eqelem\<close>, not an \<open>elem\<close>.
263\<close>
264lemmas intr_rls = NI0 NI_succ ProdI SumI PlusI_inl PlusI_inr TI
265  and intrL_rls = NI_succL ProdIL SumIL PlusI_inlL PlusI_inrL
266
267text \<open>
268  Elimination rules. OMITTED:
269  \<^item> \<open>EqE\<close>, because its conclusion is an \<open>eqelem\<close>, not an \<open>elem\<close>
270  \<^item> \<open>TE\<close>, because it does not involve a constructor.
271\<close>
272lemmas elim_rls = NE ProdE SumE PlusE FE
273  and elimL_rls = NEL ProdEL SumEL PlusEL FEL
274
275text \<open>OMITTED: \<open>eqC\<close> are \<open>TC\<close> because they make rewriting loop: \<open>p = un = un = \<dots>\<close>\<close>
276lemmas comp_rls = NC0 NC_succ ProdC SumC PlusC_inl PlusC_inr
277
278text \<open>Rules with conclusion \<open>a:A\<close>, an elem judgement.\<close>
279lemmas element_rls = intr_rls elim_rls
280
281text \<open>Definitions are (meta)equality axioms.\<close>
282lemmas basic_defs = fst_def snd_def
283
284text \<open>Compare with standard version: \<open>B\<close> is applied to UNSIMPLIFIED expression!\<close>
285lemma SumIL2: "\<lbrakk>c = a : A; d = b : B(a)\<rbrakk> \<Longrightarrow> <c,d> = <a,b> : Sum(A,B)"
286  by (rule sym_elem) (rule SumIL; rule sym_elem)
287
288lemmas intrL2_rls = NI_succL ProdIL SumIL2 PlusI_inlL PlusI_inrL
289
290text \<open>
291  Exploit \<open>p:Prod(A,B)\<close> to create the assumption \<open>z:B(a)\<close>.
292  A more natural form of product elimination.
293\<close>
294lemma subst_prodE:
295  assumes "p: Prod(A,B)"
296    and "a: A"
297    and "\<And>z. z: B(a) \<Longrightarrow> c(z): C(z)"
298  shows "c(p`a): C(p`a)"
299  by (rule assms ProdE)+
300
301
302subsection \<open>Tactics for type checking\<close>
303
304ML \<open>
305local
306
307fun is_rigid_elem (Const(\<^const_name>\<open>Elem\<close>,_) $ a $ _) = not(is_Var (head_of a))
308  | is_rigid_elem (Const(\<^const_name>\<open>Eqelem\<close>,_) $ a $ _ $ _) = not(is_Var (head_of a))
309  | is_rigid_elem (Const(\<^const_name>\<open>Type\<close>,_) $ a) = not(is_Var (head_of a))
310  | is_rigid_elem _ = false
311
312in
313
314(*Try solving a:A or a=b:A by assumption provided a is rigid!*)
315fun test_assume_tac ctxt = SUBGOAL (fn (prem, i) =>
316  if is_rigid_elem (Logic.strip_assums_concl prem)
317  then assume_tac ctxt i else no_tac)
318
319fun ASSUME ctxt tf i = test_assume_tac ctxt i  ORELSE  tf i
320
321end
322\<close>
323
324text \<open>
325  For simplification: type formation and checking,
326  but no equalities between terms.
327\<close>
328lemmas routine_rls = form_rls formL_rls refl_type element_rls
329
330ML \<open>
331fun routine_tac rls ctxt prems =
332  ASSUME ctxt (filt_resolve_from_net_tac ctxt 4 (Tactic.build_net (prems @ rls)));
333
334(*Solve all subgoals "A type" using formation rules. *)
335val form_net = Tactic.build_net @{thms form_rls};
336fun form_tac ctxt =
337  REPEAT_FIRST (ASSUME ctxt (filt_resolve_from_net_tac ctxt 1 form_net));
338
339(*Type checking: solve a:A (a rigid, A flexible) by intro and elim rules. *)
340fun typechk_tac ctxt thms =
341  let val tac =
342    filt_resolve_from_net_tac ctxt 3
343      (Tactic.build_net (thms @ @{thms form_rls} @ @{thms element_rls}))
344  in  REPEAT_FIRST (ASSUME ctxt tac)  end
345
346(*Solve a:A (a flexible, A rigid) by introduction rules.
347  Cannot use stringtrees (filt_resolve_tac) since
348  goals like ?a:SUM(A,B) have a trivial head-string *)
349fun intr_tac ctxt thms =
350  let val tac =
351    filt_resolve_from_net_tac ctxt 1
352      (Tactic.build_net (thms @ @{thms form_rls} @ @{thms intr_rls}))
353  in  REPEAT_FIRST (ASSUME ctxt tac)  end
354
355(*Equality proving: solve a=b:A (where a is rigid) by long rules. *)
356fun equal_tac ctxt thms =
357  REPEAT_FIRST
358    (ASSUME ctxt
359      (filt_resolve_from_net_tac ctxt 3
360        (Tactic.build_net (thms @ @{thms form_rls element_rls intrL_rls elimL_rls refl_elem}))))
361\<close>
362
363method_setup form = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\<close>
364method_setup typechk = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (typechk_tac ctxt ths))\<close>
365method_setup intr = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (intr_tac ctxt ths))\<close>
366method_setup equal = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (equal_tac ctxt ths))\<close>
367
368
369subsection \<open>Simplification\<close>
370
371text \<open>To simplify the type in a goal.\<close>
372lemma replace_type: "\<lbrakk>B = A; a : A\<rbrakk> \<Longrightarrow> a : B"
373  apply (rule equal_types)
374   apply (rule_tac [2] sym_type)
375   apply assumption+
376  done
377
378text \<open>Simplify the parameter of a unary type operator.\<close>
379lemma subst_eqtyparg:
380  assumes 1: "a=c : A"
381    and 2: "\<And>z. z:A \<Longrightarrow> B(z) type"
382  shows "B(a) = B(c)"
383  apply (rule subst_typeL)
384   apply (rule_tac [2] refl_type)
385   apply (rule 1)
386  apply (erule 2)
387  done
388
389text \<open>Simplification rules for Constructive Type Theory.\<close>
390lemmas reduction_rls = comp_rls [THEN trans_elem]
391
392ML \<open>
393(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification.
394  Uses other intro rules to avoid changing flexible goals.*)
395val eqintr_net = Tactic.build_net @{thms EqI intr_rls}
396fun eqintr_tac ctxt =
397  REPEAT_FIRST (ASSUME ctxt (filt_resolve_from_net_tac ctxt 1 eqintr_net))
398
399(** Tactics that instantiate CTT-rules.
400    Vars in the given terms will be incremented!
401    The (rtac EqE i) lets them apply to equality judgements. **)
402
403fun NE_tac ctxt sp i =
404  TRY (resolve_tac ctxt @{thms EqE} i) THEN
405  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm NE} i
406
407fun SumE_tac ctxt sp i =
408  TRY (resolve_tac ctxt @{thms EqE} i) THEN
409  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm SumE} i
410
411fun PlusE_tac ctxt sp i =
412  TRY (resolve_tac ctxt @{thms EqE} i) THEN
413  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm PlusE} i
414
415(** Predicate logic reasoning, WITH THINNING!!  Procedures adapted from NJ. **)
416
417(*Finds f:Prod(A,B) and a:A in the assumptions, concludes there is z:B(a) *)
418fun add_mp_tac ctxt i =
419  resolve_tac ctxt @{thms subst_prodE} i  THEN  assume_tac ctxt i  THEN  assume_tac ctxt i
420
421(*Finds P\<longrightarrow>Q and P in the assumptions, replaces implication by Q *)
422fun mp_tac ctxt i = eresolve_tac ctxt @{thms subst_prodE} i  THEN  assume_tac ctxt i
423
424(*"safe" when regarded as predicate calculus rules*)
425val safe_brls = sort (make_ord lessb)
426    [ (true, @{thm FE}), (true,asm_rl),
427      (false, @{thm ProdI}), (true, @{thm SumE}), (true, @{thm PlusE}) ]
428
429val unsafe_brls =
430    [ (false, @{thm PlusI_inl}), (false, @{thm PlusI_inr}), (false, @{thm SumI}),
431      (true, @{thm subst_prodE}) ]
432
433(*0 subgoals vs 1 or more*)
434val (safe0_brls, safep_brls) =
435    List.partition (curry (op =) 0 o subgoals_of_brl) safe_brls
436
437fun safestep_tac ctxt thms i =
438    form_tac ctxt ORELSE
439    resolve_tac ctxt thms i  ORELSE
440    biresolve_tac ctxt safe0_brls i  ORELSE  mp_tac ctxt i  ORELSE
441    DETERM (biresolve_tac ctxt safep_brls i)
442
443fun safe_tac ctxt thms i = DEPTH_SOLVE_1 (safestep_tac ctxt thms i)
444
445fun step_tac ctxt thms = safestep_tac ctxt thms  ORELSE'  biresolve_tac ctxt unsafe_brls
446
447(*Fails unless it solves the goal!*)
448fun pc_tac ctxt thms = DEPTH_SOLVE_1 o (step_tac ctxt thms)
449\<close>
450
451method_setup eqintr = \<open>Scan.succeed (SIMPLE_METHOD o eqintr_tac)\<close>
452method_setup NE = \<open>
453  Scan.lift Args.embedded_inner_syntax >> (fn s => fn ctxt => SIMPLE_METHOD' (NE_tac ctxt s))
454\<close>
455method_setup pc = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (pc_tac ctxt ths))\<close>
456method_setup add_mp = \<open>Scan.succeed (SIMPLE_METHOD' o add_mp_tac)\<close>
457
458ML_file \<open>rew.ML\<close>
459method_setup rew = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (rew_tac ctxt ths))\<close>
460method_setup hyp_rew = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_rew_tac ctxt ths))\<close>
461
462
463subsection \<open>The elimination rules for fst/snd\<close>
464
465lemma SumE_fst: "p : Sum(A,B) \<Longrightarrow> fst(p) : A"
466  apply (unfold basic_defs)
467  apply (erule SumE)
468  apply assumption
469  done
470
471text \<open>The first premise must be \<open>p:Sum(A,B)\<close>!!.\<close>
472lemma SumE_snd:
473  assumes major: "p: Sum(A,B)"
474    and "A type"
475    and "\<And>x. x:A \<Longrightarrow> B(x) type"
476  shows "snd(p) : B(fst(p))"
477  apply (unfold basic_defs)
478  apply (rule major [THEN SumE])
479  apply (rule SumC [THEN subst_eqtyparg, THEN replace_type])
480      apply (typechk assms)
481  done
482
483
484section \<open>The two-element type (booleans and conditionals)\<close>
485
486definition Bool :: "t"
487  where "Bool \<equiv> T+T"
488
489definition true :: "i"
490  where "true \<equiv> inl(tt)"
491
492definition false :: "i"
493  where "false \<equiv> inr(tt)"
494
495definition cond :: "[i,i,i]\<Rightarrow>i"
496  where "cond(a,b,c) \<equiv> when(a, \<lambda>_. b, \<lambda>_. c)"
497
498lemmas bool_defs = Bool_def true_def false_def cond_def
499
500
501subsection \<open>Derivation of rules for the type \<open>Bool\<close>\<close>
502
503text \<open>Formation rule.\<close>
504lemma boolF: "Bool type"
505  unfolding bool_defs by typechk
506
507text \<open>Introduction rules for \<open>true\<close>, \<open>false\<close>.\<close>
508
509lemma boolI_true: "true : Bool"
510  unfolding bool_defs by typechk
511
512lemma boolI_false: "false : Bool"
513  unfolding bool_defs by typechk
514
515text \<open>Elimination rule: typing of \<open>cond\<close>.\<close>
516lemma boolE: "\<lbrakk>p:Bool; a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(p,a,b) : C(p)"
517  unfolding bool_defs
518  apply (typechk; erule TE)
519   apply typechk
520  done
521
522lemma boolEL: "\<lbrakk>p = q : Bool; a = c : C(true); b = d : C(false)\<rbrakk>
523  \<Longrightarrow> cond(p,a,b) = cond(q,c,d) : C(p)"
524  unfolding bool_defs
525  apply (rule PlusEL)
526    apply (erule asm_rl refl_elem [THEN TEL])+
527  done
528
529text \<open>Computation rules for \<open>true\<close>, \<open>false\<close>.\<close>
530
531lemma boolC_true: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(true,a,b) = a : C(true)"
532  unfolding bool_defs
533  apply (rule comp_rls)
534    apply typechk
535   apply (erule_tac [!] TE)
536   apply typechk
537  done
538
539lemma boolC_false: "\<lbrakk>a : C(true); b : C(false)\<rbrakk> \<Longrightarrow> cond(false,a,b) = b : C(false)"
540  unfolding bool_defs
541  apply (rule comp_rls)
542    apply typechk
543   apply (erule_tac [!] TE)
544   apply typechk
545  done
546
547section \<open>Elementary arithmetic\<close>
548
549subsection \<open>Arithmetic operators and their definitions\<close>
550
551definition add :: "[i,i]\<Rightarrow>i"   (infixr "#+" 65)
552  where "a#+b \<equiv> rec(a, b, \<lambda>u v. succ(v))"
553
554definition diff :: "[i,i]\<Rightarrow>i"   (infixr "-" 65)
555  where "a-b \<equiv> rec(b, a, \<lambda>u v. rec(v, 0, \<lambda>x y. x))"
556
557definition absdiff :: "[i,i]\<Rightarrow>i"   (infixr "|-|" 65)
558  where "a|-|b \<equiv> (a-b) #+ (b-a)"
559
560definition mult :: "[i,i]\<Rightarrow>i"   (infixr "#*" 70)
561  where "a#*b \<equiv> rec(a, 0, \<lambda>u v. b #+ v)"
562
563definition mod :: "[i,i]\<Rightarrow>i"   (infixr "mod" 70)
564  where "a mod b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(v) |-| b, 0, \<lambda>x y. succ(v)))"
565
566definition div :: "[i,i]\<Rightarrow>i"   (infixr "div" 70)
567  where "a div b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(u) mod b, succ(v), \<lambda>x y. v))"
568
569lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def
570
571
572subsection \<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
573
574subsubsection \<open>Addition\<close>
575
576text \<open>Typing of \<open>add\<close>: short and long versions.\<close>
577
578lemma add_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #+ b : N"
579  unfolding arith_defs by typechk
580
581lemma add_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a #+ b = c #+ d : N"
582  unfolding arith_defs by equal
583
584
585text \<open>Computation for \<open>add\<close>: 0 and successor cases.\<close>
586
587lemma addC0: "b:N \<Longrightarrow> 0 #+ b = b : N"
588  unfolding arith_defs by rew
589
590lemma addC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #+ b = succ(a #+ b) : N"
591  unfolding arith_defs by rew
592
593
594subsubsection \<open>Multiplication\<close>
595
596text \<open>Typing of \<open>mult\<close>: short and long versions.\<close>
597
598lemma mult_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* b : N"
599  unfolding arith_defs by (typechk add_typing)
600
601lemma mult_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a #* b = c #* d : N"
602  unfolding arith_defs by (equal add_typingL)
603
604
605text \<open>Computation for \<open>mult\<close>: 0 and successor cases.\<close>
606
607lemma multC0: "b:N \<Longrightarrow> 0 #* b = 0 : N"
608  unfolding arith_defs by rew
609
610lemma multC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) #* b = b #+ (a #* b) : N"
611  unfolding arith_defs by rew
612
613
614subsubsection \<open>Difference\<close>
615
616text \<open>Typing of difference.\<close>
617
618lemma diff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a - b : N"
619  unfolding arith_defs by typechk
620
621lemma diff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a - b = c - d : N"
622  unfolding arith_defs by equal
623
624
625text \<open>Computation for difference: 0 and successor cases.\<close>
626
627lemma diffC0: "a:N \<Longrightarrow> a - 0 = a : N"
628  unfolding arith_defs by rew
629
630text \<open>Note: \<open>rec(a, 0, \<lambda>z w.z)\<close> is \<open>pred(a).\<close>\<close>
631
632lemma diff_0_eq_0: "b:N \<Longrightarrow> 0 - b = 0 : N"
633  unfolding arith_defs
634  apply (NE b)
635    apply hyp_rew
636  done
637
638text \<open>
639  Essential to simplify FIRST!!  (Else we get a critical pair)
640  \<open>succ(a) - succ(b)\<close> rewrites to \<open>pred(succ(a) - b)\<close>.
641\<close>
642lemma diff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) - succ(b) = a - b : N"
643  unfolding arith_defs
644  apply hyp_rew
645  apply (NE b)
646    apply hyp_rew
647  done
648
649
650subsection \<open>Simplification\<close>
651
652lemmas arith_typing_rls = add_typing mult_typing diff_typing
653  and arith_congr_rls = add_typingL mult_typingL diff_typingL
654
655lemmas congr_rls = arith_congr_rls intrL2_rls elimL_rls
656
657lemmas arithC_rls =
658  addC0 addC_succ
659  multC0 multC_succ
660  diffC0 diff_0_eq_0 diff_succ_succ
661
662ML \<open>
663  structure Arith_simp = TSimpFun(
664    val refl = @{thm refl_elem}
665    val sym = @{thm sym_elem}
666    val trans = @{thm trans_elem}
667    val refl_red = @{thm refl_red}
668    val trans_red = @{thm trans_red}
669    val red_if_equal = @{thm red_if_equal}
670    val default_rls = @{thms arithC_rls comp_rls}
671    val routine_tac = routine_tac @{thms arith_typing_rls routine_rls}
672  )
673
674  fun arith_rew_tac ctxt prems =
675    make_rew_tac ctxt (Arith_simp.norm_tac ctxt (@{thms congr_rls}, prems))
676
677  fun hyp_arith_rew_tac ctxt prems =
678    make_rew_tac ctxt
679      (Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, @{thms congr_rls}, prems))
680\<close>
681
682method_setup arith_rew = \<open>
683  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (arith_rew_tac ctxt ths))
684\<close>
685
686method_setup hyp_arith_rew = \<open>
687  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_arith_rew_tac ctxt ths))
688\<close>
689
690
691subsection \<open>Addition\<close>
692
693text \<open>Associative law for addition.\<close>
694lemma add_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #+ b) #+ c = a #+ (b #+ c) : N"
695  apply (NE a)
696    apply hyp_arith_rew
697  done
698
699text \<open>Commutative law for addition.  Can be proved using three inductions.
700  Must simplify after first induction!  Orientation of rewrites is delicate.\<close>
701lemma add_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #+ b = b #+ a : N"
702  apply (NE a)
703    apply hyp_arith_rew
704   apply (rule sym_elem)
705   prefer 2
706   apply (NE b)
707     prefer 4
708     apply (NE b)
709       apply hyp_arith_rew
710  done
711
712
713subsection \<open>Multiplication\<close>
714
715text \<open>Right annihilation in product.\<close>
716lemma mult_0_right: "a:N \<Longrightarrow> a #* 0 = 0 : N"
717  apply (NE a)
718    apply hyp_arith_rew
719  done
720
721text \<open>Right successor law for multiplication.\<close>
722lemma mult_succ_right: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* succ(b) = a #+ (a #* b) : N"
723  apply (NE a)
724    apply (hyp_arith_rew add_assoc [THEN sym_elem])
725  apply (assumption | rule add_commute mult_typingL add_typingL intrL_rls refl_elem)+
726  done
727
728text \<open>Commutative law for multiplication.\<close>
729lemma mult_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a #* b = b #* a : N"
730  apply (NE a)
731    apply (hyp_arith_rew mult_0_right mult_succ_right)
732  done
733
734text \<open>Addition distributes over multiplication.\<close>
735lemma add_mult_distrib: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #+ b) #* c = (a #* c) #+ (b #* c) : N"
736  apply (NE a)
737    apply (hyp_arith_rew add_assoc [THEN sym_elem])
738  done
739
740text \<open>Associative law for multiplication.\<close>
741lemma mult_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #* b) #* c = a #* (b #* c) : N"
742  apply (NE a)
743    apply (hyp_arith_rew add_mult_distrib)
744  done
745
746
747subsection \<open>Difference\<close>
748
749text \<open>
750  Difference on natural numbers, without negative numbers
751  \<^item> \<open>a - b = 0\<close>  iff  \<open>a \<le> b\<close>
752  \<^item> \<open>a - b = succ(c)\<close> iff \<open>a > b\<close>
753\<close>
754
755lemma diff_self_eq_0: "a:N \<Longrightarrow> a - a = 0 : N"
756  apply (NE a)
757    apply hyp_arith_rew
758  done
759
760
761lemma add_0_right: "\<lbrakk>c : N; 0 : N; c : N\<rbrakk> \<Longrightarrow> c #+ 0 = c : N"
762  by (rule addC0 [THEN [3] add_commute [THEN trans_elem]])
763
764text \<open>
765  Addition is the inverse of subtraction: if \<open>b \<le> x\<close> then \<open>b #+ (x - b) = x\<close>.
766  An example of induction over a quantified formula (a product).
767  Uses rewriting with a quantified, implicative inductive hypothesis.
768\<close>
769schematic_goal add_diff_inverse_lemma:
770  "b:N \<Longrightarrow> ?a : \<Prod>x:N. Eq(N, b-x, 0) \<longrightarrow> Eq(N, b #+ (x-b), x)"
771  apply (NE b)
772    \<comment> \<open>strip one "universal quantifier" but not the "implication"\<close>
773    apply (rule_tac [3] intr_rls)
774    \<comment> \<open>case analysis on \<open>x\<close> in \<open>succ(u) \<le> x \<longrightarrow> succ(u) #+ (x - succ(u)) = x\<close>\<close>
775     prefer 4
776     apply (NE x)
777       apply assumption
778    \<comment> \<open>Prepare for simplification of types -- the antecedent \<open>succ(u) \<le> x\<close>\<close>
779      apply (rule_tac [2] replace_type)
780       apply (rule_tac [1] replace_type)
781        apply arith_rew
782    \<comment> \<open>Solves first 0 goal, simplifies others.  Two sugbgoals remain.
783    Both follow by rewriting, (2) using quantified induction hyp.\<close>
784   apply intr \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
785    apply (hyp_arith_rew add_0_right)
786  apply assumption
787  done
788
789text \<open>
790  Version of above with premise \<open>b - a = 0\<close> i.e. \<open>a \<ge> b\<close>.
791  Using @{thm ProdE} does not work -- for \<open>?B(?a)\<close> is ambiguous.
792  Instead, @{thm add_diff_inverse_lemma} states the desired induction scheme;
793  the use of \<open>THEN\<close> below instantiates Vars in @{thm ProdE} automatically.
794\<close>
795lemma add_diff_inverse: "\<lbrakk>a:N; b:N; b - a = 0 : N\<rbrakk> \<Longrightarrow> b #+ (a-b) = a : N"
796  apply (rule EqE)
797  apply (rule add_diff_inverse_lemma [THEN ProdE, THEN ProdE])
798    apply (assumption | rule EqI)+
799  done
800
801
802subsection \<open>Absolute difference\<close>
803
804text \<open>Typing of absolute difference: short and long versions.\<close>
805
806lemma absdiff_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a |-| b : N"
807  unfolding arith_defs by typechk
808
809lemma absdiff_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a |-| b = c |-| d : N"
810  unfolding arith_defs by equal
811
812lemma absdiff_self_eq_0: "a:N \<Longrightarrow> a |-| a = 0 : N"
813  unfolding absdiff_def by (arith_rew diff_self_eq_0)
814
815lemma absdiffC0: "a:N \<Longrightarrow> 0 |-| a = a : N"
816  unfolding absdiff_def by hyp_arith_rew
817
818lemma absdiff_succ_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> succ(a) |-| succ(b)  =  a |-| b : N"
819  unfolding absdiff_def by hyp_arith_rew
820
821text \<open>Note how easy using commutative laws can be?  ...not always...\<close>
822lemma absdiff_commute: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a |-| b = b |-| a : N"
823  unfolding absdiff_def
824  apply (rule add_commute)
825   apply (typechk diff_typing)
826  done
827
828text \<open>If \<open>a + b = 0\<close> then \<open>a = 0\<close>. Surprisingly tedious.\<close>
829schematic_goal add_eq0_lemma: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> ?c : Eq(N,a#+b,0) \<longrightarrow> Eq(N,a,0)"
830  apply (NE a)
831    apply (rule_tac [3] replace_type)
832     apply arith_rew
833  apply intr  \<comment> \<open>strips remaining \<open>\<Prod>\<close>s\<close>
834   apply (rule_tac [2] zero_ne_succ [THEN FE])
835     apply (erule_tac [3] EqE [THEN sym_elem])
836    apply (typechk add_typing)
837  done
838
839text \<open>
840  Version of above with the premise \<open>a + b = 0\<close>.
841  Again, resolution instantiates variables in @{thm ProdE}.
842\<close>
843lemma add_eq0: "\<lbrakk>a:N; b:N; a #+ b = 0 : N\<rbrakk> \<Longrightarrow> a = 0 : N"
844  apply (rule EqE)
845  apply (rule add_eq0_lemma [THEN ProdE])
846    apply (rule_tac [3] EqI)
847    apply typechk
848  done
849
850text \<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>
851schematic_goal absdiff_eq0_lem:
852  "\<lbrakk>a:N; b:N; a |-| b = 0 : N\<rbrakk> \<Longrightarrow> ?a : Eq(N, a-b, 0) \<times> Eq(N, b-a, 0)"
853  apply (unfold absdiff_def)
854  apply intr
855   apply eqintr
856   apply (rule_tac [2] add_eq0)
857     apply (rule add_eq0)
858       apply (rule_tac [6] add_commute [THEN trans_elem])
859         apply (typechk diff_typing)
860  done
861
862text \<open>If \<open>a |-| b = 0\<close> then \<open>a = b\<close>
863  proof: \<open>a - b = 0\<close> and \<open>b - a = 0\<close>, so \<open>b = a + (b - a) = a + 0 = a\<close>.
864\<close>
865lemma absdiff_eq0: "\<lbrakk>a |-| b = 0 : N; a:N; b:N\<rbrakk> \<Longrightarrow> a = b : N"
866  apply (rule EqE)
867  apply (rule absdiff_eq0_lem [THEN SumE])
868     apply eqintr
869  apply (rule add_diff_inverse [THEN sym_elem, THEN trans_elem])
870     apply (erule_tac [3] EqE)
871    apply (hyp_arith_rew add_0_right)
872  done
873
874
875subsection \<open>Remainder and Quotient\<close>
876
877text \<open>Typing of remainder: short and long versions.\<close>
878
879lemma mod_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a mod b : N"
880  unfolding mod_def by (typechk absdiff_typing)
881
882lemma mod_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a mod b = c mod d : N"
883  unfolding mod_def by (equal absdiff_typingL)
884
885
886text \<open>Computation for \<open>mod\<close>: 0 and successor cases.\<close>
887
888lemma modC0: "b:N \<Longrightarrow> 0 mod b = 0 : N"
889  unfolding mod_def by (rew absdiff_typing)
890
891lemma modC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
892  succ(a) mod b = rec(succ(a mod b) |-| b, 0, \<lambda>x y. succ(a mod b)) : N"
893  unfolding mod_def by (rew absdiff_typing)
894
895
896text \<open>Typing of quotient: short and long versions.\<close>
897
898lemma div_typing: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a div b : N"
899  unfolding div_def by (typechk absdiff_typing mod_typing)
900
901lemma div_typingL: "\<lbrakk>a = c:N; b = d:N\<rbrakk> \<Longrightarrow> a div b = c div d : N"
902  unfolding div_def by (equal absdiff_typingL mod_typingL)
903
904lemmas div_typing_rls = mod_typing div_typing absdiff_typing
905
906
907text \<open>Computation for quotient: 0 and successor cases.\<close>
908
909lemma divC0: "b:N \<Longrightarrow> 0 div b = 0 : N"
910  unfolding div_def by (rew mod_typing absdiff_typing)
911
912lemma divC_succ: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
913  succ(a) div b = rec(succ(a) mod b, succ(a div b), \<lambda>x y. a div b) : N"
914  unfolding div_def by (rew mod_typing)
915
916
917text \<open>Version of above with same condition as the \<open>mod\<close> one.\<close>
918lemma divC_succ2: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow>
919  succ(a) div b =rec(succ(a mod b) |-| b, succ(a div b), \<lambda>x y. a div b) : N"
920  apply (rule divC_succ [THEN trans_elem])
921    apply (rew div_typing_rls modC_succ)
922  apply (NE "succ (a mod b) |-|b")
923    apply (rew mod_typing div_typing absdiff_typing)
924  done
925
926text \<open>For case analysis on whether a number is 0 or a successor.\<close>
927lemma iszero_decidable: "a:N \<Longrightarrow> rec(a, inl(eq), \<lambda>ka kb. inr(<ka, eq>)) :
928  Eq(N,a,0) + (\<Sum>x:N. Eq(N,a, succ(x)))"
929  apply (NE a)
930    apply (rule_tac [3] PlusI_inr)
931     apply (rule_tac [2] PlusI_inl)
932      apply eqintr
933     apply equal
934  done
935
936text \<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>
937lemma mod_div_equality: "\<lbrakk>a:N; b:N\<rbrakk> \<Longrightarrow> a mod b #+ (a div b) #* b = a : N"
938  apply (NE a)
939    apply (arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2)
940  apply (rule EqE)
941    \<comment> \<open>case analysis on \<open>succ(u mod b) |-| b\<close>\<close>
942  apply (rule_tac a1 = "succ (u mod b) |-| b" in iszero_decidable [THEN PlusE])
943    apply (erule_tac [3] SumE)
944    apply (hyp_arith_rew div_typing_rls modC0 modC_succ divC0 divC_succ2)
945    \<comment> \<open>Replace one occurrence of \<open>b\<close> by \<open>succ(u mod b)\<close>. Clumsy!\<close>
946  apply (rule add_typingL [THEN trans_elem])
947    apply (erule EqE [THEN absdiff_eq0, THEN sym_elem])
948     apply (rule_tac [3] refl_elem)
949     apply (hyp_arith_rew div_typing_rls)
950  done
951
952end
953