--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/CTT-eg.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,310 @@
+(**** CTT examples -- process using Doc/tout CTT-eg.txt ****)
+
+Pretty.setmargin 72; (*existing macros just allow this margin*)
+print_depth 0;
+
+
+(*** Type inference, from CTT/ex/typechk.ML ***)
+
+goal CTT_Rule.thy "lam n. rec(n, 0, %x y.x) : ?A";
+by (resolve_tac [ProdI] 1);
+by (eresolve_tac [NE] 2);
+by (resolve_tac [NI0] 2);
+by (assume_tac 2);
+by (resolve_tac [NF] 1);
+
+
+
+(*** Logical reasoning, from CTT/ex/elim.ML ***)
+
+val prems = goal CTT_Rule.thy
+ "[| A type; \
+\ !!x. x:A ==> B(x) type; \
+\ !!x. x:A ==> C(x) type; \
+\ p: SUM x:A. B(x) + C(x) \
+\ |] ==> ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))";
+by (resolve_tac (prems RL [SumE]) 1);
+by (eresolve_tac [PlusE] 1);
+by (resolve_tac [PlusI_inl] 1);
+by (resolve_tac [SumI] 1);
+by (assume_tac 1);
+by (assume_tac 1);
+by (typechk_tac prems);
+by (pc_tac prems 1);
+
+
+(*** Currying, from CTT/ex/elim.ML ***)
+
+val prems = goal CTT_Rule.thy
+ "[| A type; !!x. x:A ==> B(x) type; \
+\ !!z. z: (SUM x:A. B(x)) ==> C(z) type |] \
+\ ==> ?a : (PROD z : (SUM x:A . B(x)) . C(z)) \
+\ --> (PROD x:A . PROD y:B(x) . C(<x,y>))";
+by (intr_tac prems);
+by (eresolve_tac [ProdE] 1);
+by (intr_tac prems);
+
+
+(*** Axiom of Choice ***)
+
+val prems = goal CTT_Rule.thy
+ "[| A type; !!x. x:A ==> B(x) type; \
+\ !!x y.[| x:A; y:B(x) |] ==> C(x,y) type \
+\ |] ==> ?a : (PROD x:A. SUM y:B(x). C(x,y)) \
+\ --> (SUM f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))";
+by (intr_tac prems);
+by (eresolve_tac [ProdE RS SumE_fst] 1);
+by (assume_tac 1);
+by (resolve_tac [replace_type] 1);
+by (resolve_tac [subst_eqtyparg] 1);
+by (resolve_tac [ProdC] 1);
+by (typechk_tac (SumE_fst::prems));
+by (eresolve_tac [ProdE RS SumE_snd] 1);
+by (typechk_tac prems);
+
+
+
+
+> goal CTT_Rule.thy "lam n. rec(n, 0, %x y.x) : ?A";
+Level 0
+lam n. rec(n,0,%x y. x) : ?A
+ 1. lam n. rec(n,0,%x y. x) : ?A
+> by (resolve_tac [ProdI] 1);
+Level 1
+lam n. rec(n,0,%x y. x) : PROD x:?A1. ?B1(x)
+ 1. ?A1 type
+ 2. !!n. n : ?A1 ==> rec(n,0,%x y. x) : ?B1(n)
+> by (eresolve_tac [NE] 2);
+Level 2
+lam n. rec(n,0,%x y. x) : PROD x:N. ?C2(x,x)
+ 1. N type
+ 2. !!n. 0 : ?C2(n,0)
+ 3. !!n x y. [| x : N; y : ?C2(n,x) |] ==> x : ?C2(n,succ(x))
+> by (resolve_tac [NI0] 2);
+Level 3
+lam n. rec(n,0,%x y. x) : N --> N
+ 1. N type
+ 2. !!n x y. [| x : N; y : N |] ==> x : N
+> by (assume_tac 2);
+Level 4
+lam n. rec(n,0,%x y. x) : N --> N
+ 1. N type
+> by (resolve_tac [NF] 1);
+Level 5
+lam n. rec(n,0,%x y. x) : N --> N
+No subgoals!
+
+
+
+
+> val prems = goal CTT_Rule.thy
+# "[| A type; \
+# \ !!x. x:A ==> B(x) type; \
+# \ !!x. x:A ==> C(x) type; \
+# \ p: SUM x:A. B(x) + C(x) \
+# \ |] ==> ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))";
+Level 0
+?a : (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (resolve_tac (prems RL [SumE]) 1);
+Level 1
+split(p,?c1) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y.
+ [| x : A; y : B(x) + C(x) |] ==>
+ ?c1(x,y) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (eresolve_tac [PlusE] 1);
+Level 2
+split(p,%x y. when(y,?c2(x,y),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y xa.
+ [| x : A; xa : B(x) |] ==>
+ ?c2(x,y,xa) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 2. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (resolve_tac [PlusI_inl] 1);
+Level 3
+split(p,%x y. when(y,%xa. inl(?a3(x,y,xa)),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?a3(x,y,xa) : SUM x:A. B(x)
+ 2. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type
+ 3. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (resolve_tac [SumI] 1);
+Level 4
+split(p,%x y. when(y,%xa. inl(<?a4(x,y,xa),?b4(x,y,xa)>),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?a4(x,y,xa) : A
+ 2. !!x y xa. [| x : A; xa : B(x) |] ==> ?b4(x,y,xa) : B(?a4(x,y,xa))
+ 3. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type
+ 4. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (assume_tac 1);
+Level 5
+split(p,%x y. when(y,%xa. inl(<x,?b4(x,y,xa)>),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?b4(x,y,xa) : B(x)
+ 2. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type
+ 3. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (assume_tac 1);
+Level 6
+split(p,%x y. when(y,%xa. inl(<x,xa>),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type
+ 2. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (typechk_tac prems);
+Level 7
+split(p,%x y. when(y,%xa. inl(<x,xa>),?d2(x,y)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+ 1. !!x y ya.
+ [| x : A; ya : C(x) |] ==>
+ ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))
+> by (pc_tac prems 1);
+Level 8
+split(p,%x y. when(y,%xa. inl(<x,xa>),%y. inr(<x,y>)))
+: (SUM x:A. B(x)) + (SUM x:A. C(x))
+No subgoals!
+
+
+
+
+> val prems = goal CTT_Rule.thy
+# "[| A type; !!x. x:A ==> B(x) type; \
+# \ !!z. z: (SUM x:A. B(x)) ==> C(z) type |] \
+# \ ==> ?a : (PROD z : (SUM x:A . B(x)) . C(z)) \
+# \ --> (PROD x:A . PROD y:B(x) . C(<x,y>))";
+Level 0
+?a : (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))
+ 1. ?a : (PROD z:SUM x:A. B(x). C(z)) -->
+ (PROD x:A. PROD y:B(x). C(<x,y>))
+> by (intr_tac prems);
+Level 1
+lam x xa xb. ?b7(x,xa,xb)
+: (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))
+ 1. !!uu x y.
+ [| uu : PROD z:SUM x:A. B(x). C(z); x : A; y : B(x) |] ==>
+ ?b7(uu,x,y) : C(<x,y>)
+> by (eresolve_tac [ProdE] 1);
+Level 2
+lam x xa xb. x ` <xa,xb>
+: (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))
+ 1. !!uu x y. [| x : A; y : B(x) |] ==> <x,y> : SUM x:A. B(x)
+> by (intr_tac prems);
+Level 3
+lam x xa xb. x ` <xa,xb>
+: (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))
+No subgoals!
+
+
+
+
+> val prems = goal CTT_Rule.thy
+# "[| A type; !!x. x:A ==> B(x) type; \
+# \ !!x y.[| x:A; y:B(x) |] ==> C(x,y) type \
+# \ |] ==> ?a : (PROD x:A. SUM y:B(x). C(x,y)) \
+# \ --> (SUM f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))";
+Level 0
+?a : (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. ?a : (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+> by (intr_tac prems);
+Level 1
+lam x. <lam xa. ?b7(x,xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b7(uu,x) : B(x)
+ 2. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,(lam x. ?b7(uu,x)) ` x)
+> by (eresolve_tac [ProdE RS SumE_fst] 1);
+Level 2
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x. x : A ==> x : A
+ 2. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,(lam x. fst(uu ` x)) ` x)
+> by (assume_tac 1);
+Level 3
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,(lam x. fst(uu ` x)) ` x)
+> by (resolve_tac [replace_type] 1);
+Level 4
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ C(x,(lam x. fst(uu ` x)) ` x) = ?A13(uu,x)
+ 2. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : ?A13(uu,x)
+> by (resolve_tac [subst_eqtyparg] 1);
+Level 5
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ (lam x. fst(uu ` x)) ` x = ?c14(uu,x) : ?A14(uu,x)
+ 2. !!uu x z.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;
+ z : ?A14(uu,x) |] ==>
+ C(x,z) type
+ 3. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,?c14(uu,x))
+> by (resolve_tac [ProdC] 1);
+Level 6
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==> x : ?A15(uu,x)
+ 2. !!uu x xa.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;
+ xa : ?A15(uu,x) |] ==>
+ fst(uu ` xa) : ?B15(uu,x,xa)
+ 3. !!uu x z.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;
+ z : ?B15(uu,x,x) |] ==>
+ C(x,z) type
+ 4. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,fst(uu ` x))
+> by (typechk_tac (SumE_fst::prems));
+Level 7
+lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x.
+ [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>
+ ?b8(uu,x) : C(x,fst(uu ` x))
+> by (eresolve_tac [ProdE RS SumE_snd] 1);
+Level 8
+lam x. <lam xa. fst(x ` xa),lam xa. snd(x ` xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+ 1. !!uu x. x : A ==> x : A
+ 2. !!uu x. x : A ==> B(x) type
+ 3. !!uu x xa. [| x : A; xa : B(x) |] ==> C(x,xa) type
+> by (typechk_tac prems);
+Level 9
+lam x. <lam xa. fst(x ` xa),lam xa. snd(x ` xa)>
+: (PROD x:A. SUM y:B(x). C(x,y)) -->
+ (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))
+No subgoals!
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/CTT-rules.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,241 @@
+\idx{refl_type} A type ==> A = A),
+\idx{refl_elem} a : A ==> a = a : A
+
+\idx{sym_type} A = B ==> B = A
+\idx{sym_elem} a = b : A ==> b = a : A
+
+\idx{trans_type} [| A = B; B = C |] ==> A = C
+\idx{trans_elem} [| a = b : A; b = c : A |] ==> a = c : A
+
+\idx{equal_types} [| a : A; A = B |] ==> a : B
+\idx{equal_typesL} [| a = b : A; A = B |] ==> a = b : B
+
+\idx{subst_type} [| a : A; !!z. z:A ==> B(z) type |] ==> B(a) type
+
+\idx{subst_typeL} [| a = c : A; !!z. z:A ==> B(z) = D(z)
+ |] ==> B(a) = D(c)
+\idx{subst_elem} [| a : A; !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)
+
+\idx{subst_elemL} [| a = c : A; !!z. z:A ==> b(z) = d(z) : B(z)
+ |] ==> b(a) = d(c) : B(a)
+
+\idx{refl_red} Reduce(a,a)
+\idx{red_if_equal} a = b : A ==> Reduce(a,b)
+\idx{trans_red} [| a = b : A; Reduce(b,c) |] ==> a = c : A
+
+
+\idx{NF} N type
+
+\idx{NI0} 0 : N
+\idx{NI_succ} a : N ==> succ(a) : N
+\idx{NI_succL} a = b : N ==> succ(a) = succ(b) : N
+
+\idx{NE} [| p: N; a: C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u))
+ |] ==> rec(p, a, %u v.b(u,v)) : C(p)
+
+\idx{NEL} [| p = q : N; a = c : C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v) = d(u,v): C(succ(u))
+ |] ==> rec(p, a, %u v.b(u,v)) = rec(q,c,d) : C(p)
+
+\idx{NC0} [| a: C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u))
+ |] ==> rec(0, a, %u v.b(u,v)) = a : C(0)
+
+\idx{NC_succ} [| 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))
+
+\idx{zero_ne_succ} [| a: N; 0 = succ(a) : N |] ==> 0: F
+
+
+\idx{ProdF} [| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A.B(x) type
+\idx{ProdFL} [| A = C; !!x. x:A ==> B(x) = D(x) |] ==>
+ PROD x:A.B(x) = PROD x:C.D(x)
+
+\idx{ProdI} [| A type; !!x. x:A ==> b(x):B(x)
+ |] ==> lam x.b(x) : PROD x:A.B(x)
+\idx{ProdIL} [| A type; !!x. x:A ==> b(x) = c(x) : B(x)
+ |] ==> lam x.b(x) = lam x.c(x) : PROD x:A.B(x)
+
+\idx{ProdE} [| p : PROD x:A.B(x); a : A |] ==> p`a : B(a)
+\idx{ProdEL} [| p=q: PROD x:A.B(x); a=b : A |] ==> p`a = q`b : B(a)
+
+\idx{ProdC} [| a : A; !!x. x:A ==> b(x) : B(x)
+ |] ==> (lam x.b(x)) ` a = b(a) : B(a)
+
+\idx{ProdC2} p : PROD x:A.B(x) ==> (lam x. p`x) = p : PROD x:A.B(x)
+
+
+\idx{SumF} [| A type; !!x. x:A ==> B(x) type |] ==> SUM x:A.B(x) type
+\idx{SumFL} [| A = C; !!x. x:A ==> B(x) = D(x)
+ |] ==> SUM x:A.B(x) = SUM x:C.D(x)
+
+\idx{SumI} [| a : A; b : B(a) |] ==> <a,b> : SUM x:A.B(x)
+\idx{SumIL} [| a=c:A; b=d:B(a) |] ==> <a,b> = <c,d> : SUM x:A.B(x)
+
+\idx{SumE} [| p: SUM x:A.B(x);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>)
+ |] ==> split(p, %x y.c(x,y)) : C(p)
+
+\idx{SumEL} [| p=q : SUM x:A.B(x);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C(<x,y>)
+ |] ==> split(p, %x y.c(x,y)) = split(q, % x y.d(x,y)) : C(p)
+
+\idx{SumC} [| a: A; b: B(a);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>)
+ |] ==> split(<a,b>, %x y.c(x,y)) = c(a,b) : C(<a,b>)
+
+\idx{fst_def} fst(a) == split(a, %x y.x)
+\idx{snd_def} snd(a) == split(a, %x y.y)
+
+
+\idx{PlusF} [| A type; B type |] ==> A+B type
+\idx{PlusFL} [| A = C; B = D |] ==> A+B = C+D
+
+\idx{PlusI_inl} [| a : A; B type |] ==> inl(a) : A+B
+\idx{PlusI_inlL} [| a = c : A; B type |] ==> inl(a) = inl(c) : A+B
+
+\idx{PlusI_inr} [| A type; b : B |] ==> inr(b) : A+B
+\idx{PlusI_inrL} [| A type; b = d : B |] ==> inr(b) = inr(d) : A+B
+
+\idx{PlusE} [| p: A+B;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(p, %x.c(x), %y.d(y)) : C(p)
+
+\idx{PlusEL} [| p = q : A+B;
+ !!x. x: A ==> c(x) = e(x) : C(inl(x));
+ !!y. y: B ==> d(y) = f(y) : C(inr(y))
+ |] ==> when(p, %x.c(x), %y.d(y)) = when(q, %x.e(x), %y.f(y)) : C(p)
+
+\idx{PlusC_inl} [| a: A;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(inl(a), %x.c(x), %y.d(y)) = c(a) : C(inl(a))
+
+\idx{PlusC_inr} [| b: B;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(inr(b), %x.c(x), %y.d(y)) = d(b) : C(inr(b))
+
+\idx{EqF} [| A type; a : A; b : A |] ==> Eq(A,a,b) type
+\idx{EqFL} [| A=B; a=c: A; b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)
+\idx{EqI} a = b : A ==> eq : Eq(A,a,b)
+\idx{EqE} p : Eq(A,a,b) ==> a = b : A
+\idx{EqC} p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)
+
+\idx{FF} F type
+\idx{FE} [| p: F; C type |] ==> contr(p) : C
+\idx{FEL} [| p = q : F; C type |] ==> contr(p) = contr(q) : C
+
+\idx{TF} T type
+\idx{TI} tt : T
+\idx{TE} [| p : T; c : C(tt) |] ==> c : C(p)
+\idx{TEL} [| p = q : T; c = d : C(tt) |] ==> c = d : C(p)
+\idx{TC} p : T ==> p = tt : T)
+
+
+\idx{replace_type} [| B = A; a : A |] ==> a : B
+\idx{subst_eqtyparg} [| a=c : A; !!z. z:A ==> B(z) type |] ==> B(a)=B(c)
+
+\idx{subst_prodE} [| p: Prod(A,B); a: A; !!z. z: B(a) ==> c(z): C(z)
+ |] ==> c(p`a): C(p`a)
+
+\idx{SumIL2} [| c=a : A; d=b : B(a) |] ==> <c,d> = <a,b> : Sum(A,B)
+
+\idx{SumE_fst} p : Sum(A,B) ==> fst(p) : A
+
+\idx{SumE_snd} [| p: Sum(A,B); A type; !!x. x:A ==> B(x) type
+ |] ==> snd(p) : B(fst(p))
+
+
+
+\idx{add_def} a#+b == rec(a, b, %u v.succ(v))
+\idx{diff_def} a-b == rec(b, a, %u v.rec(v, 0, %x y.x))
+\idx{absdiff_def} a|-|b == (a-b) #+ (b-a)
+\idx{mult_def} a#*b == rec(a, 0, %u v. b #+ v)
+
+\idx{mod_def} a//b == rec(a, 0, %u v.
+ rec(succ(v) |-| b, 0, %x y.succ(v)))
+
+\idx{quo_def} a/b == rec(a, 0, %u v.
+ rec(succ(u) // b, succ(v), %x y.v))
+
+
+\idx{add_typing} [| a:N; b:N |] ==> a #+ b : N
+\idx{add_typingL} [| a=c:N; b=d:N |] ==> a #+ b = c #+ d : N
+\idx{addC0} b:N ==> 0 #+ b = b : N
+\idx{addC_succ} [| a:N; b:N |] ==> succ(a) #+ b = succ(a #+ b) : N
+\idx{mult_typing} [| a:N; b:N |] ==> a #* b : N
+\idx{mult_typingL} [| a=c:N; b=d:N |] ==> a #* b = c #* d : N
+\idx{multC0} b:N ==> 0 #* b = 0 : N
+\idx{multC_succ} [| a:N; b:N |] ==> succ(a) #* b = b #+ (a #* b) : N
+\idx{diff_typing} [| a:N; b:N |] ==> a - b : N
+\idx{diff_typingL} [| a=c:N; b=d:N |] ==> a - b = c - d : N
+\idx{diffC0} a:N ==> a - 0 = a : N
+\idx{diff_0_eq_0} b:N ==> 0 - b = 0 : N
+
+\idx{diff_succ_succ} [| a:N; b:N |] ==> succ(a) - succ(b) = a - b : N
+
+\idx{add_assoc} [| a:N; b:N; c:N |] ==> (a #+ b) #+ c = a #+ (b #+ c) : N
+
+\idx{add_commute} [| a:N; b:N |] ==> a #+ b = b #+ a : N
+
+\idx{mult_right0} a:N ==> a #* 0 = 0 : N
+
+\idx{mult_right_succ} [| a:N; b:N |] ==> a #* succ(b) = a #+ (a #* b) : N
+
+\idx{mult_commute} [| a:N; b:N |] ==> a #* b = b #* a : N
+
+\idx{add_mult_dist} [| a:N; b:N; c:N |] ==> (a #+ b) #* c = (a #* c) #+ (b #* c) : N
+
+\idx{mult_assoc} [| a:N; b:N; c:N |] ==> (a #* b) #* c = a #* (b #* c) : N
+
+\idx{diff_self_eq_0} a:N ==> a - a = 0 : N
+
+\idx{add_inverse_diff_lemma}
+ b:N ==> ?a : PROD x:N. Eq(N, b-x, 0) --> Eq(N, b #+ (x-b), x)
+
+\idx{add_inverse_diff} [| a:N; b:N; b-a = 0 : N |] ==> b #+ (a-b) = a : N
+
+\idx{absdiff_typing} [| a:N; b:N |] ==> a |-| b : N
+\idx{absdiff_typingL} [| a=c:N; b=d:N |] ==> a |-| b = c |-| d : N
+\idx{absdiff_self_eq_0} a:N ==> a |-| a = 0 : N
+\idx{absdiffC0} a:N ==> 0 |-| a = a : N
+
+\idx{absdiff_succ_succ} [| a:N; b:N |] ==> succ(a) |-| succ(b) = a |-| b : N
+\idx{absdiff_commute} [| a:N; b:N |] ==> a |-| b = b |-| a : N
+
+\idx{add_eq0_lemma} [| a:N; b:N |] ==> ?c : PROD u: Eq(N,a#+b,0) . Eq(N,a,0)
+
+\idx{add_eq0} [| a:N; b:N; a #+ b = 0 : N |] ==> a = 0 : N
+
+\idx{absdiff_eq0_lem}
+ [| a:N; b:N; a |-| b = 0 : N |] ==>
+ ?a : SUM v: Eq(N, a-b, 0) . Eq(N, b-a, 0)
+
+\idx{absdiff_eq0} [| a |-| b = 0 : N; a:N; b:N |] ==> a = b : N
+
+\idx{mod_typing} [| a:N; b:N |] ==> a//b : N
+\idx{mod_typingL} [| a=c:N; b=d:N |] ==> a//b = c//d : N
+\idx{modC0} b:N ==> 0//b = 0 : N
+\idx{modC_succ}
+[| a:N; b:N |] ==> succ(a)//b = rec(succ(a//b) |-| b, 0, %x y.succ(a//b)) : N
+
+\idx{quo_typing} [| a:N; b:N |] ==> a / b : N
+\idx{quo_typingL} [| a=c:N; b=d:N |] ==> a / b = c / d : N
+\idx{quoC0} b:N ==> 0 / b = 0 : N
+[| a:N; b:N |] ==> succ(a) / b =
+ rec(succ(a)//b, succ(a / b), %x y. a / b) : N
+\idx{quoC_succ2} [| a:N; b:N |] ==>
+ succ(a) / b =rec(succ(a//b) |-| b, succ(a / b), %x y. a / b) : N
+
+\idx{iszero_decidable}
+ a:N ==> rec(a, inl(eq), %ka kb.inr(<ka, eq>)) :
+ Eq(N,a,0) + (SUM x:N. Eq(N,a, succ(x)))
+
+\idx{mod_quo_equality} [| a:N; b:N |] ==> a//b #+ (a/b) #* b = a : N
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/CTT.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,1110 @@
+%% $Id$
+\chapter{Constructive Type Theory}
+Martin-L\"of's Constructive Type Theory \cite{martinlof84,nordstrom90} can
+be viewed at many different levels. It is a formal system that embodies
+the principles of intuitionistic mathematics; it embodies the
+interpretation of propositions as types; it is a vehicle for deriving
+programs from proofs. The logic is complex and many authors have attempted
+to simplify it. Thompson~\cite{thompson91} is a readable and thorough
+account of the theory.
+
+Isabelle's original formulation of Type Theory was a kind of sequent
+calculus, following Martin-L\"of~\cite{martinlof84}. It included rules for
+building the context, namely variable bindings with their types. A typical
+judgement was
+\[ a(x@1,\ldots,x@n)\in A(x@1,\ldots,x@n) \;
+ [ x@1\in A@1, x@2\in A@2(x@1), \ldots, x@n\in A@n(x@1,\ldots,x@{n-1}) ]
+\]
+This sequent calculus was not satisfactory because assumptions like
+`suppose $A$ is a type' or `suppose $B(x)$ is a type for all $x$ in $A$'
+could not be formalized.
+
+The directory~\ttindexbold{CTT} implements Constructive Type Theory, using
+natural deduction. The judgement above is expressed using $\Forall$ and
+$\Imp$:
+\[ \begin{array}{r@{}l}
+ \Forall x@1\ldots x@n. &
+ \List{x@1\in A@1;
+ x@2\in A@2(x@1); \cdots \;
+ x@n\in A@n(x@1,\ldots,x@{n-1})} \Imp \\
+ & \qquad\qquad a(x@1,\ldots,x@n)\in A(x@1,\ldots,x@n)
+ \end{array}
+\]
+Assumptions can use all the judgement forms, for instance to express that
+$B$ is a family of types over~$A$:
+\[ \Forall x . x\in A \Imp B(x)\;{\rm type} \]
+To justify the {\CTT} formulation it is probably best to appeal directly
+to the semantic explanations of the rules~\cite{martinlof84}, rather than
+to the rules themselves. The order of assumptions no longer matters,
+unlike in standard Type Theory. Contexts, which are typical of many modern
+type theories, are difficult to represent in Isabelle. In particular, it
+is difficult to enforce that all the variables in a context are distinct.
+
+The theory has the {\ML} identifier \ttindexbold{CTT.thy}. It does not
+use polymorphism. Terms in {\CTT} have type~$i$, the type of individuals.
+Types in {\CTT} have type~$t$.
+
+{\CTT} supports all of Type Theory apart from list types, well ordering
+types, and universes. Universes could be introduced {\em\`a la Tarski},
+adding new constants as names for types. The formulation {\em\`a la
+Russell}, where types denote themselves, is only possible if we identify
+the meta-types~$i$ and~$o$. Most published formulations of well ordering
+types have difficulties involving extensionality of functions; I suggest
+that you use some other method for defining recursive types. List types
+are easy to introduce by declaring new rules.
+
+{\CTT} uses the 1982 version of Type Theory, with extensional equality.
+The computation $a=b\in A$ and the equality $c\in Eq(A,a,b)$ are
+interchangeable. Its rewriting tactics prove theorems of the form $a=b\in
+A$. It could be modified to have intensional equality, but rewriting
+tactics would have to prove theorems of the form $c\in Eq(A,a,b)$ and the
+computation rules might require a second simplifier.
+
+
+\begin{figure} \tabcolsep=1em %wider spacing in tables
+\begin{center}
+\begin{tabular}{rrr}
+ \it symbol & \it meta-type & \it description \\
+ \idx{Type} & $t \to prop$ & judgement form \\
+ \idx{Eqtype} & $[t,t]\to prop$ & judgement form\\
+ \idx{Elem} & $[i, t]\to prop$ & judgement form\\
+ \idx{Eqelem} & $[i, i, t]\to prop$ & judgement form\\
+ \idx{Reduce} & $[i, i]\to prop$ & extra judgement form\\[2ex]
+
+ \idx{N} & $t$ & natural numbers type\\
+ \idx{0} & $i$ & constructor\\
+ \idx{succ} & $i\to i$ & constructor\\
+ \idx{rec} & $[i,i,[i,i]\to i]\to i$ & eliminator\\[2ex]
+ \idx{Prod} & $[t,i\to t]\to t$ & general product type\\
+ \idx{lambda} & $(i\to i)\to i$ & constructor\\[2ex]
+ \idx{Sum} & $[t, i\to t]\to t$ & general sum type\\
+ \idx{pair} & $[i,i]\to i$ & constructor\\
+ \idx{split} & $[i,[i,i]\to i]\to i$ & eliminator\\
+ \idx{fst} snd & $i\to i$ & projections\\[2ex]
+ \idx{inl} inr & $i\to i$ & constructors for $+$\\
+ \idx{when} & $[i,i\to i, i\to i]\to i$ & eliminator for $+$\\[2ex]
+ \idx{Eq} & $[t,i,i]\to t$ & equality type\\
+ \idx{eq} & $i$ & constructor\\[2ex]
+ \idx{F} & $t$ & empty type\\
+ \idx{contr} & $i\to i$ & eliminator\\[2ex]
+ \idx{T} & $t$ & singleton type\\
+ \idx{tt} & $i$ & constructor
+\end{tabular}
+\end{center}
+\caption{The constants of {\CTT}} \label{ctt-constants}
+\end{figure}
+
+
+\begin{figure} \tabcolsep=1em %wider spacing in tables
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it precedence & \it description \\
+ \idx{lam} & \idx{lambda} & $(i\To o)\To i$ & 10 & $\lambda$-abstraction
+\end{tabular}
+\end{center}
+\subcaption{Binders}
+
+\begin{center}
+\indexbold{*"`}
+\indexbold{*"+}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt ` & $[i,i]\to i$ & Left 55 & function application\\
+ \tt + & $[t,t]\to t$ & Right 30 & sum of two types
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+
+\indexbold{*"*}
+\indexbold{*"-"-">}
+\begin{center} \tt\frenchspacing
+\begin{tabular}{rrr}
+ \it external & \it internal & \it standard notation \\
+ \idx{PROD} $x$:$A$ . $B[x]$ & Prod($A$, $\lambda x.B[x]$) &
+ \rm product $\prod@{x\in A}B[x]$ \\
+ \idx{SUM} $x$:$A$ . $B[x]$ & Sum($A$, $\lambda x.B[x]$) &
+ \rm sum $\sum@{x\in A}B[x]$ \\
+ $A$ --> $B$ & Prod($A$, $\lambda x.B$) &
+ \rm function space $A\to B$ \\
+ $A$ * $B$ & Sum($A$, $\lambda x.B$) &
+ \rm binary product $A\times B$
+\end{tabular}
+\end{center}
+\subcaption{Translations}
+
+\indexbold{*"=}
+\begin{center}
+\dquotes
+\[ \begin{array}{rcl}
+prop & = & type " type" \\
+ & | & type " = " type \\
+ & | & term " : " type \\
+ & | & term " = " term " : " type
+\\[2ex]
+type & = & \hbox{expression of type~$t$} \\
+ & | & "PROD~" id " : " type " . " type \\
+ & | & "SUM~~" id " : " type " . " type
+\\[2ex]
+term & = & \hbox{expression of type~$i$} \\
+ & | & "lam " id~id^* " . " term \\
+ & | & "< " term " , " term " >"
+\end{array}
+\]
+\end{center}
+\subcaption{Grammar}
+\caption{Syntax of {\CTT}} \label{ctt-syntax}
+\end{figure}
+
+%%%%\section{Generic Packages} typedsimp.ML????????????????
+
+
+\section{Syntax}
+The constants are shown in Figure~\ref{ctt-constants}. The infixes include
+the function application operator (sometimes called `apply'), and the
+2-place type operators. Note that meta-level abstraction and application,
+$\lambda x.b$ and $f(a)$, differ from object-level abstraction and
+application, \hbox{\tt lam $x$.$b$} and $b{\tt`}a$. A {\CTT}
+function~$f$ is simply an individual as far as Isabelle is concerned: its
+Isabelle type is~$i$, not say $i\To i$.
+
+\indexbold{*F}\indexbold{*T}\indexbold{*SUM}\indexbold{*PROD}
+The empty type is called $F$ and the one-element type is $T$; other finite
+sets are built as $T+T+T$, etc. The notation for~{\CTT}
+(Figure~\ref{ctt-syntax}) is based on that of Nordstr\"om et
+al.~\cite{nordstrom90}. We can write
+\begin{ttbox}
+SUM y:B. PROD x:A. C(x,y) {\rm for} Sum(B, %y. Prod(A, %x. C(x,y)))
+\end{ttbox}
+The special cases as \hbox{\tt$A$*$B$} and \hbox{\tt$A$-->$B$} abbreviate
+general sums and products over a constant family.\footnote{Unlike normal
+infix operators, {\tt*} and {\tt-->} merely define abbreviations; there are
+no constants~{\tt op~*} and~\hbox{\tt op~-->}.} Isabelle accepts these
+abbreviations in parsing and uses them whenever possible for printing.
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{refl_type} A type ==> A = A
+\idx{refl_elem} a : A ==> a = a : A
+
+\idx{sym_type} A = B ==> B = A
+\idx{sym_elem} a = b : A ==> b = a : A
+
+\idx{trans_type} [| A = B; B = C |] ==> A = C
+\idx{trans_elem} [| a = b : A; b = c : A |] ==> a = c : A
+
+\idx{equal_types} [| a : A; A = B |] ==> a : B
+\idx{equal_typesL} [| a = b : A; A = B |] ==> a = b : B
+
+\idx{subst_type} [| a : A; !!z. z:A ==> B(z) type |] ==> B(a) type
+\idx{subst_typeL} [| a = c : A; !!z. z:A ==> B(z) = D(z)
+ |] ==> B(a) = D(c)
+
+\idx{subst_elem} [| a : A; !!z. z:A ==> b(z):B(z) |] ==> b(a):B(a)
+\idx{subst_elemL} [| a = c : A; !!z. z:A ==> b(z) = d(z) : B(z)
+ |] ==> b(a) = d(c) : B(a)
+
+\idx{refl_red} Reduce(a,a)
+\idx{red_if_equal} a = b : A ==> Reduce(a,b)
+\idx{trans_red} [| a = b : A; Reduce(b,c) |] ==> a = c : A
+\end{ttbox}
+\caption{General equality rules} \label{ctt-equality}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{NF} N type
+
+\idx{NI0} 0 : N
+\idx{NI_succ} a : N ==> succ(a) : N
+\idx{NI_succL} a = b : N ==> succ(a) = succ(b) : N
+
+\idx{NE} [| p: N; a: C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u))
+ |] ==> rec(p, a, %u v.b(u,v)) : C(p)
+
+\idx{NEL} [| p = q : N; a = c : C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v)=d(u,v): C(succ(u))
+ |] ==> rec(p, a, %u v.b(u,v)) = rec(q,c,d) : C(p)
+
+\idx{NC0} [| a: C(0);
+ !!u v. [| u: N; v: C(u) |] ==> b(u,v): C(succ(u))
+ |] ==> rec(0, a, %u v.b(u,v)) = a : C(0)
+
+\idx{NC_succ} [| 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))
+
+\idx{zero_ne_succ} [| a: N; 0 = succ(a) : N |] ==> 0: F
+\end{ttbox}
+\caption{Rules for type~$N$} \label{ctt-N}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{ProdF} [| A type; !!x. x:A ==> B(x) type |] ==> PROD x:A.B(x) type
+\idx{ProdFL} [| A = C; !!x. x:A ==> B(x) = D(x) |] ==>
+ PROD x:A.B(x) = PROD x:C.D(x)
+
+\idx{ProdI} [| A type; !!x. x:A ==> b(x):B(x)
+ |] ==> lam x.b(x) : PROD x:A.B(x)
+\idx{ProdIL} [| A type; !!x. x:A ==> b(x) = c(x) : B(x)
+ |] ==> lam x.b(x) = lam x.c(x) : PROD x:A.B(x)
+
+\idx{ProdE} [| p : PROD x:A.B(x); a : A |] ==> p`a : B(a)
+\idx{ProdEL} [| p=q: PROD x:A.B(x); a=b : A |] ==> p`a = q`b : B(a)
+
+\idx{ProdC} [| a : A; !!x. x:A ==> b(x) : B(x)
+ |] ==> (lam x.b(x)) ` a = b(a) : B(a)
+
+\idx{ProdC2} p : PROD x:A.B(x) ==> (lam x. p`x) = p : PROD x:A.B(x)
+\end{ttbox}
+\caption{Rules for the product type $\prod@{x\in A}B[x]$} \label{ctt-prod}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{SumF} [| A type; !!x. x:A ==> B(x) type |] ==> SUM x:A.B(x) type
+\idx{SumFL} [| A = C; !!x. x:A ==> B(x) = D(x)
+ |] ==> SUM x:A.B(x) = SUM x:C.D(x)
+
+\idx{SumI} [| a : A; b : B(a) |] ==> <a,b> : SUM x:A.B(x)
+\idx{SumIL} [| a=c:A; b=d:B(a) |] ==> <a,b> = <c,d> : SUM x:A.B(x)
+
+\idx{SumE} [| p: SUM x:A.B(x);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>)
+ |] ==> split(p, %x y.c(x,y)) : C(p)
+
+\idx{SumEL} [| p=q : SUM x:A.B(x);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y)=d(x,y): C(<x,y>)
+ |] ==> split(p, %x y.c(x,y)) = split(q, %x y.d(x,y)) : C(p)
+
+\idx{SumC} [| a: A; b: B(a);
+ !!x y. [| x:A; y:B(x) |] ==> c(x,y): C(<x,y>)
+ |] ==> split(<a,b>, %x y.c(x,y)) = c(a,b) : C(<a,b>)
+
+\idx{fst_def} fst(a) == split(a, %x y.x)
+\idx{snd_def} snd(a) == split(a, %x y.y)
+\end{ttbox}
+\caption{Rules for the sum type $\sum@{x\in A}B[x]$} \label{ctt-sum}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{PlusF} [| A type; B type |] ==> A+B type
+\idx{PlusFL} [| A = C; B = D |] ==> A+B = C+D
+
+\idx{PlusI_inl} [| a : A; B type |] ==> inl(a) : A+B
+\idx{PlusI_inlL} [| a = c : A; B type |] ==> inl(a) = inl(c) : A+B
+
+\idx{PlusI_inr} [| A type; b : B |] ==> inr(b) : A+B
+\idx{PlusI_inrL} [| A type; b = d : B |] ==> inr(b) = inr(d) : A+B
+
+\idx{PlusE} [| p: A+B;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(p, %x.c(x), %y.d(y)) : C(p)
+
+\idx{PlusEL} [| p = q : A+B;
+ !!x. x: A ==> c(x) = e(x) : C(inl(x));
+ !!y. y: B ==> d(y) = f(y) : C(inr(y))
+ |] ==> when(p, %x.c(x), %y.d(y)) =
+ when(q, %x.e(x), %y.f(y)) : C(p)
+
+\idx{PlusC_inl} [| a: A;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(inl(a), %x.c(x), %y.d(y)) = c(a) : C(inl(a))
+
+\idx{PlusC_inr} [| b: B;
+ !!x. x:A ==> c(x): C(inl(x));
+ !!y. y:B ==> d(y): C(inr(y))
+ |] ==> when(inr(b), %x.c(x), %y.d(y)) = d(b) : C(inr(b))
+\end{ttbox}
+\caption{Rules for the binary sum type $A+B$} \label{ctt-plus}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{EqF} [| A type; a : A; b : A |] ==> Eq(A,a,b) type
+\idx{EqFL} [| A=B; a=c: A; b=d : A |] ==> Eq(A,a,b) = Eq(B,c,d)
+\idx{EqI} a = b : A ==> eq : Eq(A,a,b)
+\idx{EqE} p : Eq(A,a,b) ==> a = b : A
+\idx{EqC} p : Eq(A,a,b) ==> p = eq : Eq(A,a,b)
+\end{ttbox}
+\subcaption{The equality type $Eq(A,a,b)$}
+
+\begin{ttbox}
+\idx{FF} F type
+\idx{FE} [| p: F; C type |] ==> contr(p) : C
+\idx{FEL} [| p = q : F; C type |] ==> contr(p) = contr(q) : C
+\end{ttbox}
+\subcaption{The empty type $F$}
+
+\begin{ttbox}
+\idx{TF} T type
+\idx{TI} tt : T
+\idx{TE} [| p : T; c : C(tt) |] ==> c : C(p)
+\idx{TEL} [| p = q : T; c = d : C(tt) |] ==> c = d : C(p)
+\idx{TC} p : T ==> p = tt : T)
+\end{ttbox}
+\subcaption{The unit type $T$}
+
+\caption{Rules for other {\CTT} types} \label{ctt-others}
+\end{figure}
+
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{replace_type} [| B = A; a : A |] ==> a : B
+\idx{subst_eqtyparg} [| a=c : A; !!z. z:A ==> B(z) type |] ==> B(a)=B(c)
+
+\idx{subst_prodE} [| p: Prod(A,B); a: A; !!z. z: B(a) ==> c(z): C(z)
+ |] ==> c(p`a): C(p`a)
+
+\idx{SumIL2} [| c=a : A; d=b : B(a) |] ==> <c,d> = <a,b> : Sum(A,B)
+
+\idx{SumE_fst} p : Sum(A,B) ==> fst(p) : A
+
+\idx{SumE_snd} [| p: Sum(A,B); A type; !!x. x:A ==> B(x) type
+ |] ==> snd(p) : B(fst(p))
+\end{ttbox}
+
+\caption{Derived rules for {\CTT}} \label{ctt-derived}
+\end{figure}
+
+
+\section{Rules of inference}
+The rules obey the following naming conventions. Type formation rules have
+the suffix~{\tt F}\@. Introduction rules have the suffix~{\tt I}\@.
+Elimination rules have the suffix~{\tt E}\@. Computation rules, which
+describe the reduction of eliminators, have the suffix~{\tt C}\@. The
+equality versions of the rules (which permit reductions on subterms) are
+called {\em long} rules; their names have the suffix~{\tt L}\@.
+Introduction and computation rules often are further suffixed with
+constructor names.
+
+Figures~\ref{ctt-equality}--\ref{ctt-others} shows the rules. Those
+for~$N$ include \ttindex{zero_ne_succ}, $0\not=n+1$: the fourth Peano axiom
+cannot be derived without universes \cite[page 91]{martinlof84}.
+Figure~\ref{ctt-sum} shows the rules for general sums, which include binary
+products as a special case, with the projections \ttindex{fst}
+and~\ttindex{snd}.
+
+The extra judgement \ttindex{Reduce} is used to implement rewriting. The
+judgement ${\tt Reduce}(a,b)$ holds when $a=b:A$ holds. It also holds
+when $a$ and $b$ are syntactically identical, even if they are ill-typed,
+because rule \ttindex{refl_red} does not verify that $a$ belongs to $A$. These
+rules do not give rise to new theorems about the standard judgements ---
+note that the only rule that makes use of {\tt Reduce} is \ttindex{trans_red},
+whose first premise ensures that $a$ and $b$ (and thus $c$) are well-typed.
+
+Derived rules are shown in Figure~\ref{ctt-derived}. The rule
+\ttindex{subst_prodE} is derived from \ttindex{prodE}, and is easier to
+use in backwards proof. The rules \ttindex{SumE_fst} and
+\ttindex{SumE_snd} express the typing of~\ttindex{fst} and~\ttindex{snd};
+together, they are roughly equivalent to~\ttindex{SumE} with the advantage
+of creating no parameters. These rules are demonstrated in a proof of the
+Axiom of Choice~(\S\ref{ctt-choice}).
+
+All the rules are given in $\eta$-expanded form. For instance, every
+occurrence of $\lambda u\,v.b(u,v)$ could be abbreviated to~$b$ in the
+rules for~$N$. This permits Isabelle to preserve bound variable names
+during backward proof. Names of bound variables in the conclusion (here,
+$u$ and~$v$) are matched with corresponding bound variables in the premises.
+
+
+\section{Rule lists}
+The Type Theory tactics provide rewriting, type inference, and logical
+reasoning. Many proof procedures work by repeatedly resolving certain Type
+Theory rules against a proof state. {\CTT} defines lists --- each with
+type
+\hbox{\tt thm list} --- of related rules.
+\begin{description}
+\item[\ttindexbold{form_rls}]
+contains formation rules for the types $N$, $\Pi$, $\Sigma$, $+$, $Eq$,
+$F$, and $T$.
+
+\item[\ttindexbold{formL_rls}]
+contains long formation rules for $\Pi$, $\Sigma$, $+$, and $Eq$. (For
+other types use \ttindex{refl_type}.)
+
+\item[\ttindexbold{intr_rls}]
+contains introduction rules for the types $N$, $\Pi$, $\Sigma$, $+$, and
+$T$.
+
+\item[\ttindexbold{intrL_rls}]
+contains long introduction rules for $N$, $\Pi$, $\Sigma$, and $+$. (For
+$T$ use \ttindex{refl_elem}.)
+
+\item[\ttindexbold{elim_rls}]
+contains elimination rules for the types $N$, $\Pi$, $\Sigma$, $+$, and
+$F$. The rules for $Eq$ and $T$ are omitted because they involve no
+eliminator.
+
+\item[\ttindexbold{elimL_rls}]
+contains long elimination rules for $N$, $\Pi$, $\Sigma$, $+$, and $F$.
+
+\item[\ttindexbold{comp_rls}]
+contains computation rules for the types $N$, $\Pi$, $\Sigma$, and $+$.
+Those for $Eq$ and $T$ involve no eliminator.
+
+\item[\ttindexbold{basic_defs}]
+contains the definitions of \ttindex{fst} and \ttindex{snd}.
+\end{description}
+
+
+\section{Tactics for subgoal reordering}
+\begin{ttbox}
+test_assume_tac : int -> tactic
+typechk_tac : thm list -> tactic
+equal_tac : thm list -> tactic
+intr_tac : thm list -> tactic
+\end{ttbox}
+Blind application of {\CTT} rules seldom leads to a proof. The elimination
+rules, especially, create subgoals containing new unknowns. These subgoals
+unify with anything, causing an undirectional search. The standard tactic
+\ttindex{filt_resolve_tac} (see the {\em Reference Manual}) can reject
+overly flexible goals; so does the {\CTT} tactic {\tt test_assume_tac}.
+Used with the tactical \ttindex{REPEAT_FIRST} they achieve a simple kind of
+subgoal reordering: the less flexible subgoals are attempted first. Do
+some single step proofs, or study the examples below, to see why this is
+necessary.
+\begin{description}
+\item[\ttindexbold{test_assume_tac} $i$]
+uses \ttindex{assume_tac} to solve the subgoal by assumption, but only if
+subgoal~$i$ has the form $a\in A$ and the head of $a$ is not an unknown.
+Otherwise, it fails.
+
+\item[\ttindexbold{typechk_tac} $thms$]
+uses $thms$ with formation, introduction, and elimination rules to check
+the typing of constructions. It is designed to solve goals of the form
+$a\in \Var{A}$, where $a$ is rigid and $\Var{A}$ is flexible. Thus it
+performs Hindley-Milner type inference. The tactic can also solve goals of
+the form $A\;\rm type$.
+
+\item[\ttindexbold{equal_tac} $thms$]
+uses $thms$ with the long introduction and elimination rules to solve goals
+of the form $a=b\in A$, where $a$ is rigid. It is intended for deriving
+the long rules for defined constants such as the arithmetic operators. The
+tactic can also perform type checking.
+
+\item[\ttindexbold{intr_tac} $thms$]
+uses $thms$ with the introduction rules to break down a type. It is
+designed for goals like $\Var{a}\in A$ where $\Var{a}$ is flexible and $A$
+rigid. These typically arise when trying to prove a proposition~$A$,
+expressed as a type.
+\end{description}
+
+
+
+\section{Rewriting tactics}
+\begin{ttbox}
+rew_tac : thm list -> tactic
+hyp_rew_tac : thm list -> tactic
+\end{ttbox}
+Object-level simplification is accomplished through proof, using the {\tt
+CTT} equality rules and the built-in rewriting functor
+\ttindex{TSimpFun}.\footnote{This should not be confused with {\tt
+SimpFun}, which is the main rewriting functor; {\tt TSimpFun} is only
+useful for {\CTT} and similar logics with type inference rules.}
+The rewrites include the computation rules and other equations. The
+long versions of the other rules permit rewriting of subterms and subtypes.
+Also used are transitivity and the extra judgement form \ttindex{Reduce}.
+Meta-level simplification handles only definitional equality.
+\begin{description}
+\item[\ttindexbold{rew_tac} $thms$]
+applies $thms$ and the computation rules as left-to-right rewrites. It
+solves the goal $a=b\in A$ by rewriting $a$ to $b$. If $b$ is an unknown
+then it is assigned the rewritten form of~$a$. All subgoals are rewritten.
+
+\item[\ttindexbold{hyp_rew_tac} $thms$]
+is like {\tt rew_tac}, but includes as rewrites any equations present in
+the assumptions.
+\end{description}
+
+
+\section{Tactics for logical reasoning}
+Interpreting propositions as types lets {\CTT} express statements of
+intuitionistic logic. However, Constructive Type Theory is not just
+another syntax for first-order logic. A key question: can assumptions be
+deleted after use? Not every occurrence of a type represents a
+proposition, and Type Theory assumptions declare variables.
+
+In first-order logic, $\disj$-elimination with the assumption $P\disj Q$
+creates one subgoal assuming $P$ and another assuming $Q$, and $P\disj Q$
+can be deleted. In Type Theory, $+$-elimination with the assumption $z\in
+A+B$ creates one subgoal assuming $x\in A$ and another assuming $y\in B$
+(for arbitrary $x$ and $y$). Deleting $z\in A+B$ may render the subgoals
+unprovable if other assumptions refer to $z$. Some people might argue that
+such subgoals are not even meaningful.
+\begin{ttbox}
+mp_tac : int -> tactic
+add_mp_tac : int -> tactic
+safestep_tac : thm list -> int -> tactic
+safe_tac : thm list -> int -> tactic
+step_tac : thm list -> int -> tactic
+pc_tac : thm list -> int -> tactic
+\end{ttbox}
+These are loosely based on the intuitionistic proof procedures
+of~\ttindex{FOL}. For the reasons discussed above, a rule that is safe for
+propositional reasoning may be unsafe for type checking; thus, some of the
+``safe'' tactics are misnamed.
+\begin{description}
+\item[\ttindexbold{mp_tac} $i$]
+searches in subgoal~$i$ for assumptions of the form $f\in\Pi(A,B)$ and
+$a\in A$, where~$A$ may be found by unification. It replaces
+$f\in\Pi(A,B)$ by $z\in B(a)$, where~$z$ is a new parameter. The tactic
+can produce multiple outcomes for each suitable pair of assumptions. In
+short, {\tt mp_tac} performs Modus Ponens among the assumptions.
+
+\item[\ttindexbold{add_mp_tac} $i$]
+is like {\tt mp_tac}~$i$ but retains the assumption $f\in\Pi(A,B)$.
+
+\item[\ttindexbold{safestep_tac} $thms$ $i$]
+attacks subgoal~$i$ using formation rules and certain other `safe' rules
+(\ttindex{FE}, \ttindex{ProdI}, \ttindex{SumE}, \ttindex{PlusE}), calling
+{\tt mp_tac} when appropriate. It also uses~$thms$,
+which are typically premises of the rule being derived.
+
+\item[\ttindexbold{safe_tac} $thms$ $i$]
+tries to solve subgoal~$i$ by backtracking, using {\tt safestep_tac}.
+
+\item[\ttindexbold{step_tac} $thms$ $i$]
+tries to reduce subgoal~$i$ using {\tt safestep_tac}, then tries unsafe
+rules. It may produce multiple outcomes.
+
+\item[\ttindexbold{pc_tac} $thms$ $i$]
+tries to solve subgoal~$i$ by backtracking, using {\tt step_tac}.
+\end{description}
+
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{add_def} a#+b == rec(a, b, %u v.succ(v))
+\idx{diff_def} a-b == rec(b, a, %u v.rec(v, 0, %x y.x))
+\idx{absdiff_def} a|-|b == (a-b) #+ (b-a)
+\idx{mult_def} a#*b == rec(a, 0, %u v. b #+ v)
+
+\idx{mod_def} a//b == rec(a, 0,
+ %u v. rec(succ(v) |-| b, 0, %x y.succ(v)))
+
+\idx{quo_def} a/b == rec(a, 0,
+ %u v. rec(succ(u) // b, succ(v), %x y.v))
+\end{ttbox}
+\subcaption{Definitions of the operators}
+
+\begin{ttbox}
+\idx{add_typing} [| a:N; b:N |] ==> a #+ b : N
+\idx{addC0} b:N ==> 0 #+ b = b : N
+\idx{addC_succ} [| a:N; b:N |] ==> succ(a) #+ b = succ(a #+ b) : N
+
+\idx{add_assoc} [| a:N; b:N; c:N |] ==>
+ (a #+ b) #+ c = a #+ (b #+ c) : N
+
+\idx{add_commute} [| a:N; b:N |] ==> a #+ b = b #+ a : N
+
+\idx{mult_typing} [| a:N; b:N |] ==> a #* b : N
+\idx{multC0} b:N ==> 0 #* b = 0 : N
+\idx{multC_succ} [| a:N; b:N |] ==> succ(a) #* b = b #+ (a#*b) : N
+\idx{mult_commute} [| a:N; b:N |] ==> a #* b = b #* a : N
+
+\idx{add_mult_dist} [| a:N; b:N; c:N |] ==>
+ (a #+ b) #* c = (a #* c) #+ (b #* c) : N
+
+\idx{mult_assoc} [| a:N; b:N; c:N |] ==>
+ (a #* b) #* c = a #* (b #* c) : N
+
+\idx{diff_typing} [| a:N; b:N |] ==> a - b : N
+\idx{diffC0} a:N ==> a - 0 = a : N
+\idx{diff_0_eq_0} b:N ==> 0 - b = 0 : N
+\idx{diff_succ_succ} [| a:N; b:N |] ==> succ(a) - succ(b) = a - b : N
+\idx{diff_self_eq_0} a:N ==> a - a = 0 : N
+\idx{add_inverse_diff} [| a:N; b:N; b-a=0 : N |] ==> b #+ (a-b) = a : N
+\end{ttbox}
+\subcaption{Some theorems of arithmetic}
+\caption{The theory of arithmetic} \label{ctt-arith}
+\end{figure}
+
+
+\section{A theory of arithmetic}
+{\CTT} contains a theory of elementary arithmetic. It proves the
+properties of addition, multiplication, subtraction, division, and
+remainder, culminating in the theorem
+\[ a \bmod b + (a/b)\times b = a. \]
+Figure~\ref{ctt-arith} presents the definitions and some of the key
+theorems, including commutative, distributive, and associative laws. The
+theory has the {\ML} identifier \ttindexbold{arith.thy}. All proofs are on
+the file \ttindexbold{CTT/arith.ML}.
+
+The operators~\verb'#+', \verb'-', \verb'|-|', \verb'#*', \verb'//'
+and~\verb'/' stand for sum, difference, absolute difference, product,
+remainder and quotient, respectively. Since Type Theory has only primitive
+recursion, some of their definitions may be obscure.
+
+The difference~$a-b$ is computed by taking $b$ predecessors of~$a$, where
+the predecessor function is $\lambda v. {\tt rec}(v, 0, \lambda x\,y.x)$.
+
+The remainder $a//b$ counts up to~$a$ in a cyclic fashion, using 0 as the
+successor of~$b-1$. Absolute difference is used to test the equality
+$succ(v)=b$.
+
+The quotient $a//b$ is computed by adding one for every number $x$ such
+that $0\leq x \leq a$ and $x//b = 0$.
+
+
+
+\section{The examples directory}
+This directory contains examples and experimental proofs in {\CTT}.
+\begin{description}
+\item[\ttindexbold{CTT/ex/typechk.ML}]
+contains simple examples of type checking and type deduction.
+
+\item[\ttindexbold{CTT/ex/elim.ML}]
+contains some examples from Martin-L\"of~\cite{martinlof84}, proved using
+{\tt pc_tac}.
+
+\item[\ttindexbold{CTT/ex/equal.ML}]
+contains simple examples of rewriting.
+
+\item[\ttindexbold{CTT/ex/synth.ML}]
+demonstrates the use of unknowns with some trivial examples of program
+synthesis.
+\end{description}
+
+
+\section{Example: type inference}
+Type inference involves proving a goal of the form $a\in\Var{A}$, where $a$
+is a term and $\Var{A}$ is an unknown standing for its type. The type,
+initially
+unknown, takes shape in the course of the proof. Our example is the
+predecessor function on the natural numbers.
+\begin{ttbox}
+goal CTT.thy "lam n. rec(n, 0, %x y.x) : ?A";
+{\out Level 0}
+{\out lam n. rec(n,0,%x y. x) : ?A}
+{\out 1. lam n. rec(n,0,%x y. x) : ?A}
+\end{ttbox}
+Since the term is a Constructive Type Theory $\lambda$-abstraction (not to
+be confused with a meta-level abstraction), we apply the rule
+\ttindex{ProdI}, for $\Pi$-introduction. This instantiates~$\Var{A}$ to a
+product type of unknown domain and range.
+\begin{ttbox}
+by (resolve_tac [ProdI] 1);
+{\out Level 1}
+{\out lam n. rec(n,0,%x y. x) : PROD x:?A1. ?B1(x)}
+{\out 1. ?A1 type}
+{\out 2. !!n. n : ?A1 ==> rec(n,0,%x y. x) : ?B1(n)}
+\end{ttbox}
+Subgoal~1 can be solved by instantiating~$\Var{A@1}$ to any type, but this
+could invalidate subgoal~2. We therefore tackle the latter subgoal. It
+asks the type of a term beginning with {\tt rec}, which can be found by
+$N$-elimination.\index{*NE}
+\begin{ttbox}
+by (eresolve_tac [NE] 2);
+{\out Level 2}
+{\out lam n. rec(n,0,%x y. x) : PROD x:N. ?C2(x,x)}
+{\out 1. N type}
+{\out 2. !!n. 0 : ?C2(n,0)}
+{\out 3. !!n x y. [| x : N; y : ?C2(n,x) |] ==> x : ?C2(n,succ(x))}
+\end{ttbox}
+We now know~$\Var{A@1}$ is the type of natural numbers. However, let us
+continue with subgoal~2. What is the type of~0?\index{*NIO}
+\begin{ttbox}
+by (resolve_tac [NI0] 2);
+{\out Level 3}
+{\out lam n. rec(n,0,%x y. x) : N --> N}
+{\out 1. N type}
+{\out 2. !!n x y. [| x : N; y : N |] ==> x : N}
+\end{ttbox}
+The type~$\Var{A}$ is now determined. It is $\prod@{n\in N}N$, which is
+equivalent to $N\to N$. But we must prove all the subgoals to show that
+the original term is validly typed. Subgoal~2 is provable by assumption
+and the remaining subgoal falls by $N$-formation.\index{*NF}
+\begin{ttbox}
+by (assume_tac 2);
+{\out Level 4}
+{\out lam n. rec(n,0,%x y. x) : N --> N}
+{\out 1. N type}
+by (resolve_tac [NF] 1);
+{\out Level 5}
+{\out lam n. rec(n,0,%x y. x) : N --> N}
+{\out No subgoals!}
+\end{ttbox}
+Calling \ttindex{typechk_tac} can prove this theorem in one step.
+
+
+\section{An example of logical reasoning}
+Logical reasoning in Type Theory involves proving a goal of the form
+$\Var{a}\in A$, where type $A$ expresses a proposition and $\Var{a}$ is an
+unknown standing
+for its proof term: a value of type $A$. This term is initially unknown, as
+with type inference, and takes shape during the proof. Our example
+expresses, by propositions-as-types, a theorem about quantifiers in a
+sorted logic:
+\[ \infer{(\ex{x\in A}P(x)) \disj (\ex{x\in A}Q(x))}
+ {\ex{x\in A}P(x)\disj Q(x)}
+\]
+It it related to a distributive law of Type Theory:
+\[ \infer{(A\times B) + (A\times C)}{A\times(B+C)} \]
+Generalizing this from $\times$ to $\Sigma$, and making the typing
+conditions explicit, yields
+\[ \infer{\Var{a} \in (\sum@{x\in A} B(x)) + (\sum@{x\in A} C(x))}
+ {\hbox{$A$ type} &
+ \infer*{\hbox{$B(x)$ type}}{[x\in A]} &
+ \infer*{\hbox{$C(x)$ type}}{[x\in A]} &
+ p\in \sum@{x\in A} B(x)+C(x)}
+\]
+To derive this rule, we bind its premises --- returned by~\ttindex{goal}
+--- to the {\ML} variable~{\tt prems}.
+\begin{ttbox}
+val prems = goal CTT.thy
+ "[| A type; \ttback
+\ttback !!x. x:A ==> B(x) type; \ttback
+\ttback !!x. x:A ==> C(x) type; \ttback
+\ttback p: SUM x:A. B(x) + C(x) \ttback
+\ttback |] ==> ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))";
+{\out Level 0}
+{\out ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. ?a : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+One of the premises involves summation ($\Sigma$). Since it is a premise
+rather than the assumption of a goal, it cannot be found by
+\ttindex{eresolve_tac}. We could insert it by calling
+\hbox{\tt \ttindex{cut_facts_tac} prems 1}. Instead, let us resolve the
+$\Sigma$-elimination rule with the premises; this yields one result, which
+we supply to \ttindex{resolve_tac}.\index{*SumE}\index{*RL}
+\begin{ttbox}
+by (resolve_tac (prems RL [SumE]) 1);
+{\out Level 1}
+{\out split(p,?c1) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y.}
+{\out [| x : A; y : B(x) + C(x) |] ==>}
+{\out ?c1(x,y) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+The subgoal has two new parameters. In the main goal, $\Var{a}$ has been
+instantiated with a \ttindex{split} term. The assumption $y\in B(x) + C(x)$ is
+eliminated next, causing a case split and a new parameter. The main goal
+now contains~\ttindex{when}.
+\index{*PlusE}
+\begin{ttbox}
+by (eresolve_tac [PlusE] 1);
+{\out Level 2}
+{\out split(p,%x y. when(y,?c2(x,y),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y xa.}
+{\out [| x : A; xa : B(x) |] ==>}
+{\out ?c2(x,y,xa) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 2. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+To complete the proof object for the main goal, we need to instantiate the
+terms $\Var{c@2}(x,y,xa)$ and $\Var{d@2}(x,y,xa)$. We attack subgoal~1 by
+introduction of~$+$; since it assumes $xa\in B(x)$, we take the left
+injection~(\ttindex{inl}).
+\index{*PlusI_inl}
+\begin{ttbox}
+by (resolve_tac [PlusI_inl] 1);
+{\out Level 3}
+{\out split(p,%x y. when(y,%xa. inl(?a3(x,y,xa)),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?a3(x,y,xa) : SUM x:A. B(x)}
+{\out 2. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type}
+{\out 3. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+A new subgoal has appeared, to verify that $\sum@{x\in A}C(x)$ is a type.
+Continuing with subgoal~1, we apply $\Sigma$-introduction. The main goal
+now contains an ordered pair.
+\index{*SumI}
+\begin{ttbox}
+by (resolve_tac [SumI] 1);
+{\out Level 4}
+{\out split(p,%x y. when(y,%xa. inl(<?a4(x,y,xa),?b4(x,y,xa)>),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?a4(x,y,xa) : A}
+{\out 2. !!x y xa. [| x : A; xa : B(x) |] ==> ?b4(x,y,xa) : B(?a4(x,y,xa))}
+{\out 3. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type}
+{\out 4. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+The two new subgoals both hold by assumption. Observe how the unknowns
+$\Var{a@4}$ and $\Var{b@4}$ are instantiated throughout the proof state.
+\begin{ttbox}
+by (assume_tac 1);
+{\out Level 5}
+{\out split(p,%x y. when(y,%xa. inl(<x,?b4(x,y,xa)>),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y xa. [| x : A; xa : B(x) |] ==> ?b4(x,y,xa) : B(x)}
+{\out 2. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type}
+{\out 3. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+by (assume_tac 1);
+{\out Level 6}
+{\out split(p,%x y. when(y,%xa. inl(<x,xa>),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y xa. [| x : A; xa : B(x) |] ==> SUM x:A. C(x) type}
+{\out 2. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+Subgoal~1 is just type checking. It yields to \ttindex{typechk_tac},
+supplied with the current list of premises.
+\begin{ttbox}
+by (typechk_tac prems);
+{\out Level 7}
+{\out split(p,%x y. when(y,%xa. inl(<x,xa>),?d2(x,y)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out 1. !!x y ya.}
+{\out [| x : A; ya : C(x) |] ==>}
+{\out ?d2(x,y,ya) : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+\end{ttbox}
+The other case is similar. Let us prove it by \ttindex{pc_tac}, and note
+the final proof object.
+\begin{ttbox}
+by (pc_tac prems 1);
+{\out Level 8}
+{\out split(p,%x y. when(y,%xa. inl(<x,xa>),%y. inr(<x,y>)))}
+{\out : (SUM x:A. B(x)) + (SUM x:A. C(x))}
+{\out No subgoals!}
+\end{ttbox}
+Calling \ttindex{pc_tac} after the first $\Sigma$-elimination above also
+proves this theorem.
+
+
+\section{Example: deriving a currying functional}
+In simply-typed languages such as {\ML}, a currying functional has the type
+\[ (A\times B \to C) \to (A\to (B\to C)). \]
+Let us generalize this to~$\Sigma$ and~$\Pi$. The argument of the
+functional is a function that maps $z:\Sigma(A,B)$ to~$C(z)$; the resulting
+function maps $x\in A$ and $y\in B(x)$ to $C(\langle x,y\rangle)$. Here
+$B$ is a family over~$A$, while $C$ is a family over $\Sigma(A,B)$.
+\begin{ttbox}
+val prems = goal CTT.thy
+ "[| A type; !!x. x:A ==> B(x) type; \ttback
+\ttback !!z. z: (SUM x:A. B(x)) ==> C(z) type |] \ttback
+\ttback ==> ?a : (PROD z : (SUM x:A . B(x)) . C(z)) \ttback
+\ttback --> (PROD x:A . PROD y:B(x) . C(<x,y>))";
+{\out Level 0}
+{\out ?a : (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))}
+{\out 1. ?a : (PROD z:SUM x:A. B(x). C(z)) -->}
+{\out (PROD x:A. PROD y:B(x). C(<x,y>))}
+\end{ttbox}
+This is an opportunity to demonstrate \ttindex{intr_tac}. Here, the tactic
+repeatedly applies $\Pi$-introduction, automatically proving the rather
+tiresome typing conditions. Note that $\Var{a}$ becomes instantiated to
+three nested $\lambda$-abstractions.
+\begin{ttbox}
+by (intr_tac prems);
+{\out Level 1}
+{\out lam x xa xb. ?b7(x,xa,xb)}
+{\out : (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))}
+{\out 1. !!uu x y.}
+{\out [| uu : PROD z:SUM x:A. B(x). C(z); x : A; y : B(x) |] ==>}
+{\out ?b7(uu,x,y) : C(<x,y>)}
+\end{ttbox}
+Using $\Pi$-elimination, we solve subgoal~1 by applying the function~$uu$.
+\index{*ProdE}
+\begin{ttbox}
+by (eresolve_tac [ProdE] 1);
+{\out Level 2}
+{\out lam x xa xb. x ` <xa,xb>}
+{\out : (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))}
+{\out 1. !!uu x y. [| x : A; y : B(x) |] ==> <x,y> : SUM x:A. B(x)}
+\end{ttbox}
+Finally, we exhibit a suitable argument for the function application. This
+is straightforward using introduction rules.
+\index{*intr_tac}
+\begin{ttbox}
+by (intr_tac prems);
+{\out Level 3}
+{\out lam x xa xb. x ` <xa,xb>}
+{\out : (PROD z:SUM x:A. B(x). C(z)) --> (PROD x:A. PROD y:B(x). C(<x,y>))}
+{\out No subgoals!}
+\end{ttbox}
+Calling~\ttindex{pc_tac} would have proved this theorem in one step; it can
+also prove an example by Martin-L\"of, related to $\disj$-elimination
+\cite[page~58]{martinlof84}.
+
+
+\section{Example: proving the Axiom of Choice} \label{ctt-choice}
+Suppose we have a function $h\in \prod@{x\in A}\sum@{y\in B(x)} C(x,y)$,
+which takes $x\in A$ to some $y\in B(x)$ paired with some $z\in C(x,y)$.
+Interpreting propositions as types, this asserts that for all $x\in A$
+there exists $y\in B(x)$ such that $C(x,y)$. The Axiom of Choice asserts
+that we can construct a function $f\in \prod@{x\in A}B(x)$ such that
+$C(x,f{\tt`}x)$ for all $x\in A$, where the latter property is witnessed by a
+function $g\in \prod@{x\in A}C(x,f{\tt`}x)$.
+
+In principle, the Axiom of Choice is simple to derive in Constructive Type
+Theory \cite[page~50]{martinlof84}. The following definitions work:
+\begin{eqnarray*}
+ f & \equiv & {\tt fst} \circ h \\
+ g & \equiv & {\tt snd} \circ h
+\end{eqnarray*}
+But a completely formal proof is hard to find. Many of the rules can be
+applied in a multiplicity of ways, yielding a large number of higher-order
+unifiers. The proof can get bogged down in the details. But with a
+careful selection of derived rules (recall Figure~\ref{ctt-derived}) and
+the type checking tactics, we can prove the theorem in nine steps.
+\begin{ttbox}
+val prems = goal CTT.thy
+ "[| A type; !!x. x:A ==> B(x) type; \ttback
+\ttback !!x y.[| x:A; y:B(x) |] ==> C(x,y) type \ttback
+\ttback |] ==> ?a : (PROD x:A. SUM y:B(x). C(x,y)) \ttback
+\ttback --> (SUM f: (PROD x:A. B(x)). PROD x:A. C(x, f`x))";
+{\out Level 0}
+{\out ?a : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. ?a : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+\end{ttbox}
+First, \ttindex{intr_tac} applies introduction rules and performs routine
+type checking. This instantiates~$\Var{a}$ to a construction involving
+three $\lambda$-abstractions and an ordered pair.
+\begin{ttbox}
+by (intr_tac prems);
+{\out Level 1}
+{\out lam x. <lam xa. ?b7(x,xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b7(uu,x) : B(x)}
+{\out 2. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,(lam x. ?b7(uu,x)) ` x)}
+\end{ttbox}
+Subgoal~1 asks to find the choice function itself, taking $x\in A$ to some
+$\Var{b@7}(uu,x)\in B(x)$. Subgoal~2 asks, given $x\in A$, for a proof
+object $\Var{b@8}(uu,x)$ to witness that the choice function's argument
+and result lie in the relation~$C$.
+\index{*ProdE}\index{*SumE_fst}\index{*RS}
+\begin{ttbox}
+by (eresolve_tac [ProdE RS SumE_fst] 1);
+{\out Level 2}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x. x : A ==> x : A}
+{\out 2. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,(lam x. fst(uu ` x)) ` x)}
+\end{ttbox}
+Above, we have composed \ttindex{fst} with the function~$h$ (named~$uu$ in
+the assumptions). Unification has deduced that the function must be
+applied to $x\in A$.
+\begin{ttbox}
+by (assume_tac 1);
+{\out Level 3}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,(lam x. fst(uu ` x)) ` x)}
+\end{ttbox}
+Before we can compose \ttindex{snd} with~$h$, the arguments of $C$ must be
+simplified. The derived rule \ttindex{replace_type} lets us replace a type
+by any equivalent type:
+\begin{ttbox}
+by (resolve_tac [replace_type] 1);
+{\out Level 4}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out C(x,(lam x. fst(uu ` x)) ` x) = ?A13(uu,x)}
+{\out 2. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : ?A13(uu,x)}
+\end{ttbox}
+The derived rule \ttindex{subst_eqtyparg} lets us simplify a type's
+argument (by currying, $C(x)$ is a unary type operator):
+\begin{ttbox}
+by (resolve_tac [subst_eqtyparg] 1);
+{\out Level 5}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out (lam x. fst(uu ` x)) ` x = ?c14(uu,x) : ?A14(uu,x)}
+{\out 2. !!uu x z.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;}
+{\out z : ?A14(uu,x) |] ==>}
+{\out C(x,z) type}
+{\out 3. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,?c14(uu,x))}
+\end{ttbox}
+The rule \ttindex{ProdC} is simply $\beta$-reduction. The term
+$\Var{c@{14}}(uu,x)$ receives the simplified form, $f{\tt`}x$.
+\begin{ttbox}
+by (resolve_tac [ProdC] 1);
+{\out Level 6}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==> x : ?A15(uu,x)}
+{\out 2. !!uu x xa.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;}
+{\out xa : ?A15(uu,x) |] ==>}
+{\out fst(uu ` xa) : ?B15(uu,x,xa)}
+{\out 3. !!uu x z.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A;}
+{\out z : ?B15(uu,x,x) |] ==>}
+{\out C(x,z) type}
+{\out 4. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,fst(uu ` x))}
+\end{ttbox}
+Routine type checking goals proliferate in Constructive Type Theory, but
+\ttindex{typechk_tac} quickly solves them. Note the inclusion of
+\ttindex{SumE_fst}.
+\begin{ttbox}
+by (typechk_tac (SumE_fst::prems));
+{\out Level 7}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. ?b8(x,xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x.}
+{\out [| uu : PROD x:A. SUM y:B(x). C(x,y); x : A |] ==>}
+{\out ?b8(uu,x) : C(x,fst(uu ` x))}
+\end{ttbox}
+We are finally ready to compose \ttindex{snd} with~$h$.
+\index{*ProdE}\index{*SumE_snd}\index{*RS}
+\begin{ttbox}
+by (eresolve_tac [ProdE RS SumE_snd] 1);
+{\out Level 8}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. snd(x ` xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out 1. !!uu x. x : A ==> x : A}
+{\out 2. !!uu x. x : A ==> B(x) type}
+{\out 3. !!uu x xa. [| x : A; xa : B(x) |] ==> C(x,xa) type}
+\end{ttbox}
+The proof object has reached its final form. We call \ttindex{typechk_tac}
+to finish the type checking.
+\begin{ttbox}
+by (typechk_tac prems);
+{\out Level 9}
+{\out lam x. <lam xa. fst(x ` xa),lam xa. snd(x ` xa)>}
+{\out : (PROD x:A. SUM y:B(x). C(x,y)) -->}
+{\out (SUM f:PROD x:A. B(x). PROD x:A. C(x,f ` x))}
+{\out No subgoals!}
+\end{ttbox}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/FOL-eg.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,317 @@
+(**** FOL examples ****)
+
+Pretty.setmargin 72; (*existing macros just allow this margin*)
+print_depth 0;
+
+(*** Intuitionistic examples ***)
+
+(*Quantifier example from Logic&Computation*)
+goal Int_Rule.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
+by (resolve_tac [impI] 1);
+by (resolve_tac [allI] 1);
+by (resolve_tac [exI] 1);
+by (eresolve_tac [exE] 1);
+choplev 2;
+by (eresolve_tac [exE] 1);
+by (resolve_tac [exI] 1);
+by (eresolve_tac [allE] 1);
+by (assume_tac 1);
+
+
+(*Example of Dyckhoff's method*)
+goalw Int_Rule.thy [not_def] "~ ~ ((P-->Q) | (Q-->P))";
+by (resolve_tac [impI] 1);
+by (eresolve_tac [disj_impE] 1);
+by (eresolve_tac [imp_impE] 1);
+by (eresolve_tac [imp_impE] 1);
+by (REPEAT (eresolve_tac [FalseE] 2));
+by (assume_tac 1);
+
+
+
+- goal Int_Rule.thy "(EX ay. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
+Level 0
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+- by (resolve_tac [impI] 1);
+Level 1
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. EX y. ALL x. Q(x,y) ==> ALL x. EX y. Q(x,y)
+- by (resolve_tac [allI] 1);
+Level 2
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)
+- by (resolve_tac [exI] 1);
+Level 3
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x. EX y. ALL x. Q(x,y) ==> Q(x,?y2(x))
+- by (eresolve_tac [exE] 1);
+Level 4
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y2(x))
+- choplev 2;
+Level 2
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)
+- by (eresolve_tac [exE] 1);
+Level 3
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x y. ALL x. Q(x,y) ==> EX y. Q(x,y)
+- by (resolve_tac [exI] 1);
+Level 4
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y3(x,y))
+- by (eresolve_tac [allE] 1);
+Level 5
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+ 1. !!x y. Q(?x4(x,y),y) ==> Q(x,?y3(x,y))
+- by (assume_tac 1);
+Level 6
+(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))
+No subgoals!
+
+
+
+> goalw Int_Rule.thy [not_def] "~ ~ ((P-->Q) | (Q-->P))";
+Level 0
+~ ~ ((P --> Q) | (Q --> P))
+ 1. ((P --> Q) | (Q --> P) --> False) --> False
+> by (resolve_tac [impI] 1);
+Level 1
+~ ~ ((P --> Q) | (Q --> P))
+ 1. (P --> Q) | (Q --> P) --> False ==> False
+> by (eresolve_tac [disj_impE] 1);
+Level 2
+~ ~ ((P --> Q) | (Q --> P))
+ 1. [| (P --> Q) --> False; (Q --> P) --> False |] ==> False
+> by (eresolve_tac [imp_impE] 1);
+Level 3
+~ ~ ((P --> Q) | (Q --> P))
+ 1. [| (Q --> P) --> False; P; Q --> False |] ==> Q
+ 2. [| (Q --> P) --> False; False |] ==> False
+> by (eresolve_tac [imp_impE] 1);
+Level 4
+~ ~ ((P --> Q) | (Q --> P))
+ 1. [| P; Q --> False; Q; P --> False |] ==> P
+ 2. [| P; Q --> False; False |] ==> Q
+ 3. [| (Q --> P) --> False; False |] ==> False
+> by (REPEAT (eresolve_tac [FalseE] 2));
+Level 5
+~ ~ ((P --> Q) | (Q --> P))
+ 1. [| P; Q --> False; Q; P --> False |] ==> P
+> by (assume_tac 1);
+Level 6
+~ ~ ((P --> Q) | (Q --> P))
+No subgoals!
+
+
+
+
+(*** Classical examples ***)
+
+goal cla_thy "EX y. ALL x. P(y)-->P(x)";
+by (resolve_tac [exCI] 1);
+by (resolve_tac [allI] 1);
+by (resolve_tac [impI] 1);
+by (eresolve_tac [allE] 1);
+prth (allI RSN (2,swap));
+by (eresolve_tac [it] 1);
+by (resolve_tac [impI] 1);
+by (eresolve_tac [notE] 1);
+by (assume_tac 1);
+goal cla_thy "EX y. ALL x. P(y)-->P(x)";
+by (best_tac FOL_dup_cs 1);
+
+
+
+- goal cla_thy "EX y. ALL x. P(y)-->P(x)";
+Level 0
+EX y. ALL x. P(y) --> P(x)
+ 1. EX y. ALL x. P(y) --> P(x)
+- by (resolve_tac [exCI] 1);
+Level 1
+EX y. ALL x. P(y) --> P(x)
+ 1. ALL y. ~(ALL x. P(y) --> P(x)) ==> ALL x. P(?a) --> P(x)
+- by (resolve_tac [allI] 1);
+Level 2
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x. ALL y. ~(ALL x. P(y) --> P(x)) ==> P(?a) --> P(x)
+- by (resolve_tac [impI] 1);
+Level 3
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x. [| ALL y. ~(ALL x. P(y) --> P(x)); P(?a) |] ==> P(x)
+- by (eresolve_tac [allE] 1);
+Level 4
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x. [| P(?a); ~(ALL xa. P(?y3(x)) --> P(xa)) |] ==> P(x)
+- prth (allI RSN (2,swap));
+[| ~(ALL x. ?P1(x)); !!x. ~?Q ==> ?P1(x) |] ==> ?Q
+- by (eresolve_tac [it] 1);
+Level 5
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x xa. [| P(?a); ~P(x) |] ==> P(?y3(x)) --> P(xa)
+- by (resolve_tac [impI] 1);
+Level 6
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x xa. [| P(?a); ~P(x); P(?y3(x)) |] ==> P(xa)
+- by (eresolve_tac [notE] 1);
+Level 7
+EX y. ALL x. P(y) --> P(x)
+ 1. !!x xa. [| P(?a); P(?y3(x)) |] ==> P(x)
+- by (assume_tac 1);
+Level 8
+EX y. ALL x. P(y) --> P(x)
+No subgoals!
+- goal cla_thy "EX y. ALL x. P(y)-->P(x)";
+Level 0
+EX y. ALL x. P(y) --> P(x)
+ 1. EX y. ALL x. P(y) --> P(x)
+- by (best_tac FOL_dup_cs 1);
+Level 1
+EX y. ALL x. P(y) --> P(x)
+No subgoals!
+
+
+(**** finally, the example FOL/ex/if.ML ****)
+
+> val prems = goalw if_thy [if_def]
+# "[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)";
+Level 0
+if(P,Q,R)
+ 1. P & Q | ~P & R
+> by (Classical.fast_tac (FOL_cs addIs prems) 1);
+Level 1
+if(P,Q,R)
+No subgoals!
+> val ifI = result();
+
+
+> val major::prems = goalw if_thy [if_def]
+# "[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S";
+Level 0
+S
+ 1. S
+> by (cut_facts_tac [major] 1);
+Level 1
+S
+ 1. P & Q | ~P & R ==> S
+> by (Classical.fast_tac (FOL_cs addIs prems) 1);
+Level 2
+S
+No subgoals!
+> val ifE = result();
+
+> goal if_thy "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))";
+Level 0
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+> by (resolve_tac [iffI] 1);
+Level 1
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. if(P,if(Q,A,B),if(Q,C,D)) ==> if(Q,if(P,A,C),if(P,B,D))
+ 2. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))
+> by (eresolve_tac [ifE] 1);
+Level 2
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. [| P; if(Q,A,B) |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 2. [| ~P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 3. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))
+> by (eresolve_tac [ifE] 1);
+Level 3
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. [| P; Q; A |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 2. [| P; ~Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 3. [| ~P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 4. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))
+> by (resolve_tac [ifI] 1);
+Level 4
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. [| P; Q; A; Q |] ==> if(P,A,C)
+ 2. [| P; Q; A; ~Q |] ==> if(P,B,D)
+ 3. [| P; ~Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 4. [| ~P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 5. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))
+> by (resolve_tac [ifI] 1);
+Level 5
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. [| P; Q; A; Q; P |] ==> A
+ 2. [| P; Q; A; Q; ~P |] ==> C
+ 3. [| P; Q; A; ~Q |] ==> if(P,B,D)
+ 4. [| P; ~Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 5. [| ~P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))
+ 6. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))
+
+> choplev 0;
+Level 0
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+ 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+> val if_cs = FOL_cs addSIs [ifI] addSEs[ifE];
+> by (Classical.fast_tac if_cs 1);
+Level 1
+if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))
+No subgoals!
+> val if_commute = result();
+
+> goal if_thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))";
+Level 0
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+ 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+> by (Classical.fast_tac if_cs 1);
+Level 1
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+No subgoals!
+> val nested_ifs = result();
+
+
+> choplev 0;
+Level 0
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+ 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+> by (rewrite_goals_tac [if_def]);
+Level 1
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+ 1. (P & Q | ~P & R) & A | ~(P & Q | ~P & R) & B <->
+ P & (Q & A | ~Q & B) | ~P & (R & A | ~R & B)
+> by (Classical.fast_tac FOL_cs 1);
+Level 2
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))
+No subgoals!
+
+
+> goal if_thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))";
+Level 0
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+ 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+> by (REPEAT (Classical.step_tac if_cs 1));
+Level 1
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+ 1. [| A; ~P; R; ~P; R |] ==> B
+ 2. [| B; ~P; ~R; ~P; ~R |] ==> A
+ 3. [| ~P; R; B; ~P; R |] ==> A
+ 4. [| ~P; ~R; A; ~B; ~P |] ==> R
+
+> choplev 0;
+Level 0
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+ 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+> by (rewrite_goals_tac [if_def]);
+Level 1
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+ 1. (P & Q | ~P & R) & A | ~(P & Q | ~P & R) & B <->
+ P & (Q & A | ~Q & B) | ~P & (R & B | ~R & A)
+> by (Classical.fast_tac FOL_cs 1);
+by: tactic failed
+Exception- ERROR raised
+Exception failure raised
+
+> by (REPEAT (Classical.step_tac FOL_cs 1));
+Level 2
+if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))
+ 1. [| A; ~P; R; ~P; R; ~False |] ==> B
+ 2. [| A; ~P; R; R; ~False; ~B; ~B |] ==> Q
+ 3. [| B; ~P; ~R; ~P; ~A |] ==> R
+ 4. [| B; ~P; ~R; ~Q; ~A |] ==> R
+ 5. [| B; ~R; ~P; ~A; ~R; Q; ~False |] ==> A
+ 6. [| ~P; R; B; ~P; R; ~False |] ==> A
+ 7. [| ~P; ~R; A; ~B; ~R |] ==> P
+ 8. [| ~P; ~R; A; ~B; ~R |] ==> Q
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/FOL.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,918 @@
+%% $Id$
+\chapter{First-order logic}
+The directory~\ttindexbold{FOL} contains theories for first-order logic
+based on Gentzen's natural deduction systems (which he called {\sc nj} and
+{\sc nk}). Intuitionistic logic is defined first; then classical logic is
+obtained by adding the double negation rule. Basic proof procedures are
+provided. The intuitionistic prover works with derived rules to simplify
+implications in the assumptions. Classical logic makes use of Isabelle's
+generic prover for classical reasoning, which simulates a sequent calculus.
+
+\section{Syntax and rules of inference}
+The logic is many-sorted, using Isabelle's type classes. The
+class of first-order terms is called {\it term} and is a subclass of
+{\it logic}. No types of individuals
+are provided, but extensions can define types such as $nat::term$ and type
+constructors such as $list::(term)term$. See the examples directory.
+Below, the type variable $\alpha$ ranges over class {\it term\/}; the
+equality symbol and quantifiers are polymorphic (many-sorted). The type of
+formulae is~{\it o}, which belongs to class {\it logic}.
+Figure~\ref{fol-syntax} gives the syntax. Note that $a$\verb|~=|$b$ is
+translated to \verb|~(|$a$=$b$\verb|)|.
+
+The intuitionistic theory has the \ML\ identifier
+\ttindexbold{IFOL.thy}. Figure~\ref{fol-rules} shows the inference
+rules with their~\ML\ names. Negation is defined in the usual way for
+intuitionistic logic; $\neg P$ abbreviates $P\imp\bot$. The
+biconditional~($\bimp$) is defined through $\conj$ and~$\imp$; introduction
+and elimination rules are derived for it.
+
+The unique existence quantifier, $\exists!x.P(x)$, is defined in terms
+of~$\exists$ and~$\forall$. An Isabelle binder, it admits nested
+quantifications. For instance, $\exists!x y.P(x,y)$ abbreviates
+$\exists!x. \exists!y.P(x,y)$; note that this does not mean that there
+exists a unique pair $(x,y)$ satisfying~$P(x,y)$.
+
+Some intuitionistic derived rules are shown in
+Figure~\ref{fol-int-derived}, again with their \ML\ names. These include
+rules for the defined symbols $\neg$, $\bimp$ and $\exists!$. Natural
+deduction typically involves a combination of forwards and backwards
+reasoning, particularly with the destruction rules $(\conj E)$,
+$({\imp}E)$, and~$(\forall E)$. Isabelle's backwards style handles these
+rules badly, so sequent-style rules are derived to eliminate conjunctions,
+implications, and universal quantifiers. Used with elim-resolution,
+\ttindex{allE} eliminates a universal quantifier while \ttindex{all_dupE}
+re-inserts the quantified formula for later use. The rules {\tt
+conj_impE}, etc., support the intuitionistic proof procedure
+(see~\S\ref{fol-int-prover}).
+
+See the files \ttindexbold{FOL/ifol.thy}, \ttindexbold{FOL/ifol.ML} and
+\ttindexbold{FOL/int-prover.ML} for complete listings of the rules and
+derived rules.
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{Trueprop}& $o\To prop$ & coercion to $prop$\\
+ \idx{Not} & $o\To o$ & negation ($\neg$) \\
+ \idx{True} & $o$ & tautology ($\top$) \\
+ \idx{False} & $o$ & absurdity ($\bot$)
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it precedence & \it description \\
+ \idx{ALL} & \idx{All} & $(\alpha\To o)\To o$ & 10 &
+ universal quantifier ($\forall$) \\
+ \idx{EX} & \idx{Ex} & $(\alpha\To o)\To o$ & 10 &
+ existential quantifier ($\exists$) \\
+ \idx{EX!} & \idx{Ex1} & $(\alpha\To o)\To o$ & 10 &
+ unique existence ($\exists!$)
+\end{tabular}
+\end{center}
+\subcaption{Binders}
+
+\begin{center}
+\indexbold{*"=}
+\indexbold{&@{\tt\&}}
+\indexbold{*"|}
+\indexbold{*"-"-">}
+\indexbold{*"<"-">}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt = & $[\alpha,\alpha]\To o$ & Left 50 & equality ($=$) \\
+ \tt \& & $[o,o]\To o$ & Right 35 & conjunction ($\conj$) \\
+ \tt | & $[o,o]\To o$ & Right 30 & disjunction ($\disj$) \\
+ \tt --> & $[o,o]\To o$ & Right 25 & implication ($\imp$) \\
+ \tt <-> & $[o,o]\To o$ & Right 25 & biconditional ($\bimp$)
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+
+\dquotes
+\[\begin{array}{rcl}
+ formula & = & \hbox{expression of type~$o$} \\
+ & | & term " = " term \\
+ & | & term " \ttilde= " term \\
+ & | & "\ttilde\ " formula \\
+ & | & formula " \& " formula \\
+ & | & formula " | " formula \\
+ & | & formula " --> " formula \\
+ & | & formula " <-> " formula \\
+ & | & "ALL~" id~id^* " . " formula \\
+ & | & "EX~~" id~id^* " . " formula \\
+ & | & "EX!~" id~id^* " . " formula
+ \end{array}
+\]
+\subcaption{Grammar}
+\caption{Syntax of {\tt FOL}} \label{fol-syntax}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{refl} a=a
+\idx{subst} [| a=b; P(a) |] ==> P(b)
+\subcaption{Equality rules}
+
+\idx{conjI} [| P; Q |] ==> P&Q
+\idx{conjunct1} P&Q ==> P
+\idx{conjunct2} P&Q ==> Q
+
+\idx{disjI1} P ==> P|Q
+\idx{disjI2} Q ==> P|Q
+\idx{disjE} [| P|Q; P ==> R; Q ==> R |] ==> R
+
+\idx{impI} (P ==> Q) ==> P-->Q
+\idx{mp} [| P-->Q; P |] ==> Q
+
+\idx{FalseE} False ==> P
+\subcaption{Propositional rules}
+
+\idx{allI} (!!x. P(x)) ==> (ALL x.P(x))
+\idx{spec} (ALL x.P(x)) ==> P(x)
+
+\idx{exI} P(x) ==> (EX x.P(x))
+\idx{exE} [| EX x.P(x); !!x. P(x) ==> R |] ==> R
+\subcaption{Quantifier rules}
+
+\idx{True_def} True == False-->False
+\idx{not_def} ~P == P-->False
+\idx{iff_def} P<->Q == (P-->Q) & (Q-->P)
+\idx{ex1_def} EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)
+\subcaption{Definitions}
+\end{ttbox}
+
+\caption{Rules of intuitionistic {\tt FOL}} \label{fol-rules}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{sym} a=b ==> b=a
+\idx{trans} [| a=b; b=c |] ==> a=c
+\idx{ssubst} [| b=a; P(a) |] ==> P(b)
+\subcaption{Derived equality rules}
+
+\idx{TrueI} True
+
+\idx{notI} (P ==> False) ==> ~P
+\idx{notE} [| ~P; P |] ==> R
+
+\idx{iffI} [| P ==> Q; Q ==> P |] ==> P<->Q
+\idx{iffE} [| P <-> Q; [| P-->Q; Q-->P |] ==> R |] ==> R
+\idx{iffD1} [| P <-> Q; P |] ==> Q
+\idx{iffD2} [| P <-> Q; Q |] ==> P
+
+\idx{ex1I} [| P(a); !!x. P(x) ==> x=a |] ==> EX! x. P(x)
+\idx{ex1E} [| EX! x.P(x); !!x.[| P(x); ALL y. P(y) --> y=x |] ==> R
+ |] ==> R
+\subcaption{Derived rules for \(\top\), \(\neg\), \(\bimp\) and \(\exists!\)}
+
+\idx{conjE} [| P&Q; [| P; Q |] ==> R |] ==> R
+\idx{impE} [| P-->Q; P; Q ==> R |] ==> R
+\idx{allE} [| ALL x.P(x); P(x) ==> R |] ==> R
+\idx{all_dupE} [| ALL x.P(x); [| P(x); ALL x.P(x) |] ==> R |] ==> R
+\subcaption{Sequent-style elimination rules}
+
+\idx{conj_impE} [| (P&Q)-->S; P-->(Q-->S) ==> R |] ==> R
+\idx{disj_impE} [| (P|Q)-->S; [| P-->S; Q-->S |] ==> R |] ==> R
+\idx{imp_impE} [| (P-->Q)-->S; [| P; Q-->S |] ==> Q; S ==> R |] ==> R
+\idx{not_impE} [| ~P --> S; P ==> False; S ==> R |] ==> R
+\idx{iff_impE} [| (P<->Q)-->S; [| P; Q-->S |] ==> Q; [| Q; P-->S |] ==> P;
+ S ==> R |] ==> R
+\idx{all_impE} [| (ALL x.P(x))-->S; !!x.P(x); S ==> R |] ==> R
+\idx{ex_impE} [| (EX x.P(x))-->S; P(a)-->S ==> R |] ==> R
+\end{ttbox}
+\subcaption{Intuitionistic simplification of implication}
+\caption{Derived rules for {\tt FOL}} \label{fol-int-derived}
+\end{figure}
+
+
+\section{Generic packages}
+\FOL{} instantiates most of Isabelle's generic packages;
+see \ttindexbold{FOL/ROOT.ML} for details.
+\begin{itemize}
+\item
+Because it includes a general substitution rule, \FOL{} instantiates the
+tactic \ttindex{hyp_subst_tac}, which substitutes for an equality
+throughout a subgoal and its hypotheses.
+\item
+It instantiates the simplifier. \ttindexbold{IFOL_ss} is the simplification
+set for intuitionistic first-order logic, while \ttindexbold{FOL_ss} is the
+simplification set for classical logic. Both equality ($=$) and the
+biconditional ($\bimp$) may be used for rewriting. See the file
+\ttindexbold{FOL/simpdata.ML} for a complete listing of the simplification
+rules.
+\item
+It instantiates the classical reasoning module. See~\S\ref{fol-cla-prover}
+for details.
+\end{itemize}
+
+
+\section{Intuitionistic proof procedures} \label{fol-int-prover}
+Implication elimination (the rules~{\tt mp} and~{\tt impE}) pose
+difficulties for automated proof. Given $P\imp Q$ we may assume $Q$
+provided we can prove $P$. In classical logic the proof of $P$ can assume
+$\neg P$, but the intuitionistic proof of $P$ may require repeated use of
+$P\imp Q$. If the proof of $P$ fails then the whole branch of the proof
+must be abandoned. Thus intuitionistic propositional logic requires
+backtracking. For an elementary example, consider the intuitionistic proof
+of $Q$ from $P\imp Q$ and $(P\imp Q)\imp P$. The implication $P\imp Q$ is
+needed twice:
+\[ \infer[({\imp}E)]{Q}{P\imp Q &
+ \infer[({\imp}E)]{P}{(P\imp Q)\imp P & P\imp Q}}
+\]
+The theorem prover for intuitionistic logic does not use~{\tt impE}.\@
+Instead, it simplifies implications using derived rules
+(Figure~\ref{fol-int-derived}). It reduces the antecedents of implications
+to atoms and then uses Modus Ponens: from $P\imp Q$ and $P$ deduce~$Q$.
+The rules \ttindex{conj_impE} and \ttindex{disj_impE} are
+straightforward: $(P\conj Q)\imp S$ is equivalent to $P\imp (Q\imp S)$, and
+$(P\disj Q)\imp S$ is equivalent to the conjunction of $P\imp S$ and $Q\imp
+S$. The other \ldots{\tt_impE} rules are unsafe; the method requires
+backtracking. Observe that \ttindex{imp_impE} does not admit the (unsound)
+inference of~$P$ from $(P\imp Q)\imp S$. All the rules are derived in
+essentially the same simple manner.
+
+Dyckhoff has independently discovered similar rules, and (more importantly)
+has demonstrated their completeness for propositional
+logic~\cite{dyckhoff}. However, the tactics given below are not complete
+for first-order logic because they discard universally quantified
+assumptions after a single use.
+\begin{ttbox}
+mp_tac : int -> tactic
+eq_mp_tac : int -> tactic
+Int.safe_step_tac : int -> tactic
+Int.safe_tac : tactic
+Int.step_tac : int -> tactic
+Int.fast_tac : int -> tactic
+Int.best_tac : int -> tactic
+\end{ttbox}
+Most of these belong to the structure \ttindexbold{Int}. They are
+similar or identical to tactics (with the same names) provided by
+Isabelle's classical module (see {\em The Isabelle Reference Manual\/}).
+
+\begin{description}
+\item[\ttindexbold{mp_tac} {\it i}]
+attempts to use \ttindex{notE} or \ttindex{impE} within the assumptions in
+subgoal $i$. For each assumption of the form $\neg P$ or $P\imp Q$, it
+searches for another assumption unifiable with~$P$. By
+contradiction with $\neg P$ it can solve the subgoal completely; by Modus
+Ponens it can replace the assumption $P\imp Q$ by $Q$. The tactic can
+produce multiple outcomes, enumerating all suitable pairs of assumptions.
+
+\item[\ttindexbold{eq_mp_tac} {\it i}]
+is like {\tt mp_tac} {\it i}, but may not instantiate unknowns --- thus, it
+is safe.
+
+\item[\ttindexbold{Int.safe_step_tac} $i$] performs a safe step on
+subgoal~$i$. This may include proof by assumption or Modus Ponens, taking
+care not to instantiate unknowns, or \ttindex{hyp_subst_tac}.
+
+\item[\ttindexbold{Int.safe_tac}] repeatedly performs safe steps on all
+subgoals. It is deterministic, with at most one outcome.
+
+\item[\ttindexbold{Int.inst_step_tac} $i$] is like {\tt safe_step_tac},
+but allows unknowns to be instantiated.
+
+\item[\ttindexbold{step_tac} $i$] tries {\tt safe_tac} or {\tt
+inst_step_tac}, or applies an unsafe rule. This is the basic step of the
+proof procedure.
+
+\item[\ttindexbold{Int.step_tac} $i$] tries {\tt safe_tac} or
+certain unsafe inferences. This is the basic step of the intuitionistic
+proof procedure.
+
+\item[\ttindexbold{Int.fast_tac} $i$] applies {\tt step_tac}, using
+depth-first search, to solve subgoal~$i$.
+
+\item[\ttindexbold{Int.best_tac} $i$] applies {\tt step_tac}, using
+best-first search (guided by the size of the proof state) to solve subgoal~$i$.
+\end{description}
+Here are some of the theorems that {\tt Int.fast_tac} proves
+automatically. The latter three date from {\it Principia Mathematica}
+(*11.53, *11.55, *11.61)~\cite{principia}.
+\begin{ttbox}
+~~P & ~~(P --> Q) --> ~~Q
+(ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))
+(EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))
+(EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))
+\end{ttbox}
+
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{excluded_middle} ~P | P
+
+\idx{disjCI} (~Q ==> P) ==> P|Q
+\idx{exCI} (ALL x. ~P(x) ==> P(a)) ==> EX x.P(x)
+\idx{impCE} [| P-->Q; ~P ==> R; Q ==> R |] ==> R
+\idx{iffCE} [| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R
+\idx{notnotD} ~~P ==> P
+\idx{swap} ~P ==> (~Q ==> P) ==> Q
+\end{ttbox}
+\caption{Derived rules for classical logic} \label{fol-cla-derived}
+\end{figure}
+
+
+\section{Classical proof procedures} \label{fol-cla-prover}
+The classical theory has the \ML\ identifier \ttindexbold{FOL.thy}. It
+consists of intuitionistic logic plus the rule
+$$ \vcenter{\infer{P}{\infer*{P}{[\neg P]}}} \eqno(classical) $$
+\noindent
+Natural deduction in classical logic is not really all that natural.
+{\FOL} derives classical introduction rules for $\disj$ and~$\exists$, as
+well as classical elimination rules for~$\imp$ and~$\bimp$, and the swap
+rule (see Figure~\ref{fol-cla-derived}).
+
+The classical reasoning module is set up for \FOL, as the
+structure~\ttindexbold{Cla}. This structure is open, so \ML{} identifiers
+such as {\tt step_tac}, {\tt fast_tac}, {\tt best_tac}, etc., refer to it.
+Single-step proofs can be performed, using \ttindex{swap_res_tac} to deal
+with negated assumptions.
+
+{\FOL} defines the following classical rule sets:
+\begin{ttbox}
+prop_cs : claset
+FOL_cs : claset
+FOL_dup_cs : claset
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{prop_cs}] contains the propositional rules, namely
+those for~$\top$, $\bot$, $\conj$, $\disj$, $\neg$, $\imp$ and~$\bimp$,
+along with the rule~\ttindex{refl}.
+
+\item[\ttindexbold{FOL_cs}]
+extends {\tt prop_cs} with the safe rules \ttindex{allI} and~\ttindex{exE}
+and the unsafe rules \ttindex{allE} and~\ttindex{exI}, as well as rules for
+unique existence. Search using this is incomplete since quantified
+formulae are used at most once.
+
+\item[\ttindexbold{FOL_dup_cs}]
+extends {\tt prop_cs} with the safe rules \ttindex{allI} and~\ttindex{exE}
+and the unsafe rules \ttindex{all_dupE} and~\ttindex{exCI}, as well as
+rules for unique existence. Search using this is complete --- quantified
+formulae may be duplicated --- but frequently fails to terminate. It is
+generally unsuitable for depth-first search.
+\end{description}
+\noindent
+See the file \ttindexbold{FOL/fol.ML} for derivations of the
+classical rules, and the {\em Reference Manual} for more discussion of
+classical proof methods.
+
+
+\section{An intuitionistic example}
+Here is a session similar to one in {\em Logic and Computation}
+\cite[pages~222--3]{paulson87}. Isabelle treats quantifiers differently
+from {\sc lcf}-based theorem provers such as {\sc hol}. The proof begins
+by entering the goal in intuitionistic logic, then applying the rule
+$({\imp}I)$.
+\begin{ttbox}
+goal IFOL.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
+{\out Level 0}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+\ttbreak
+by (resolve_tac [impI] 1);
+{\out Level 1}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. EX y. ALL x. Q(x,y) ==> ALL x. EX y. Q(x,y)}
+\end{ttbox}
+In this example we will never have more than one subgoal. Applying
+$({\imp}I)$ replaces~\verb|-->| by~\verb|==>|, making
+\(\ex{y}\all{x}Q(x,y)\) an assumption. We have the choice of
+$({\exists}E)$ and $({\forall}I)$; let us try the latter.
+\begin{ttbox}
+by (resolve_tac [allI] 1);
+{\out Level 2}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
+\end{ttbox}
+Applying $({\forall}I)$ replaces the \hbox{\tt ALL x} by \hbox{\tt!!x},
+changing the universal quantifier from object~($\forall$) to
+meta~($\Forall$). The bound variable is a {\em parameter\/} of the
+subgoal. We now must choose between $({\exists}I)$ and $({\exists}E)$. What
+happens if the wrong rule is chosen?
+\begin{ttbox}
+by (resolve_tac [exI] 1);
+{\out Level 3}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x. EX y. ALL x. Q(x,y) ==> Q(x,?y2(x))}
+\end{ttbox}
+The new subgoal~1 contains the function variable {\tt?y2}. Instantiating
+{\tt?y2} can replace~{\tt?y2(x)} by a term containing~{\tt x}, even
+though~{\tt x} is a bound variable. Now we analyse the assumption
+\(\exists y.\forall x. Q(x,y)\) using elimination rules:
+\begin{ttbox}
+by (eresolve_tac [exE] 1);
+{\out Level 4}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y2(x))}
+\end{ttbox}
+Applying $(\exists E)$ has produced the parameter {\tt y} and stripped the
+existential quantifier from the assumption. But the subgoal is unprovable.
+There is no way to unify {\tt ?y2(x)} with the bound variable~{\tt y}:
+assigning \verb|%x.y| to {\tt ?y2} is illegal because {\tt ?y2} is in the
+scope of the bound variable {\tt y}. Using \ttindex{choplev} we
+can return to the mistake. This time we apply $({\exists}E)$:
+\begin{ttbox}
+choplev 2;
+{\out Level 2}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x. EX y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
+\ttbreak
+by (eresolve_tac [exE] 1);
+{\out Level 3}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x y. ALL x. Q(x,y) ==> EX y. Q(x,y)}
+\end{ttbox}
+We now have two parameters and no scheme variables. Parameters should be
+produced early. Applying $({\exists}I)$ and $({\forall}E)$ will produce
+two scheme variables.
+\begin{ttbox}
+by (resolve_tac [exI] 1);
+{\out Level 4}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x y. ALL x. Q(x,y) ==> Q(x,?y3(x,y))}
+\ttbreak
+by (eresolve_tac [allE] 1);
+{\out Level 5}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. !!x y. Q(?x4(x,y),y) ==> Q(x,?y3(x,y))}
+\end{ttbox}
+The subgoal has variables {\tt ?y3} and {\tt ?x4} applied to both
+parameters. The obvious projection functions unify {\tt?x4(x,y)} with~{\tt
+x} and \verb|?y3(x,y)| with~{\tt y}.
+\begin{ttbox}
+by (assume_tac 1);
+{\out Level 6}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out No subgoals!}
+\end{ttbox}
+The theorem was proved in six tactic steps, not counting the abandoned
+ones. But proof checking is tedious; {\tt Int.fast_tac} proves the
+theorem in one step.
+\begin{ttbox}
+goal IFOL.thy "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))";
+{\out Level 0}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out 1. (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+by (Int.fast_tac 1);
+{\out Level 1}
+{\out (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\section{An example of intuitionistic negation}
+The following example demonstrates the specialized forms of implication
+elimination. Even propositional formulae can be difficult to prove from
+the basic rules; the specialized rules help considerably.
+
+Propositional examples are easy to invent, for as Dummett notes~\cite[page
+28]{dummett}, $\neg P$ is classically provable if and only if it is
+intuitionistically provable. Therefore, $P$ is classically provable if and
+only if $\neg\neg P$ is intuitionistically provable. In both cases, $P$
+must be a propositional formula (no quantifiers). Our example,
+accordingly, is the double negation of a classical tautology: $(P\imp
+Q)\disj (Q\imp P)$.
+
+When stating the goal, we command Isabelle to expand the negation symbol,
+using the definition $\neg P\equiv P\imp\bot$. Although negation possesses
+derived rules that effect precisely this definition --- the automatic
+tactics apply them --- it seems more straightforward to unfold the
+negations.
+\begin{ttbox}
+goalw IFOL.thy [not_def] "~ ~ ((P-->Q) | (Q-->P))";
+{\out Level 0}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. ((P --> Q) | (Q --> P) --> False) --> False}
+\end{ttbox}
+The first step is trivial.
+\begin{ttbox}
+by (resolve_tac [impI] 1);
+{\out Level 1}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. (P --> Q) | (Q --> P) --> False ==> False}
+\end{ttbox}
+Proving $(P\imp Q)\disj (Q\imp P)$ would suffice, but this formula is
+constructively invalid. Instead we apply \ttindex{disj_impE}. It splits
+the assumption into two parts, one for each disjunct.
+\begin{ttbox}
+by (eresolve_tac [disj_impE] 1);
+{\out Level 2}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. [| (P --> Q) --> False; (Q --> P) --> False |] ==> False}
+\end{ttbox}
+We cannot hope to prove $P\imp Q$ or $Q\imp P$ separately, but
+their negations are inconsistent. Applying \ttindex{imp_impE} breaks down
+the assumption $\neg(P\imp Q)$, asking to show~$Q$ while providing new
+assumptions~$P$ and~$\neg Q$.
+\begin{ttbox}
+by (eresolve_tac [imp_impE] 1);
+{\out Level 3}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. [| (Q --> P) --> False; P; Q --> False |] ==> Q}
+{\out 2. [| (Q --> P) --> False; False |] ==> False}
+\end{ttbox}
+Subgoal~2 holds trivially; let us ignore it and continue working on
+subgoal~1. Thanks to the assumption~$P$, we could prove $Q\imp P$;
+applying \ttindex{imp_impE} is simpler.
+\begin{ttbox}
+by (eresolve_tac [imp_impE] 1);
+{\out Level 4}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. [| P; Q --> False; Q; P --> False |] ==> P}
+{\out 2. [| P; Q --> False; False |] ==> Q}
+{\out 3. [| (Q --> P) --> False; False |] ==> False}
+\end{ttbox}
+The three subgoals are all trivial.
+\begin{ttbox}
+by (REPEAT (eresolve_tac [FalseE] 2));
+{\out Level 5}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out 1. [| P; Q --> False; Q; P --> False |] ==> P}
+by (assume_tac 1);
+{\out Level 6}
+{\out ~ ~ ((P --> Q) | (Q --> P))}
+{\out No subgoals!}
+\end{ttbox}
+This proof is also trivial for {\tt Int.fast_tac}.
+
+
+\section{A classical example} \label{fol-cla-example}
+To illustrate classical logic, we shall prove the theorem
+$\ex{y}\all{x}P(y)\imp P(x)$. Classically, the theorem can be proved as
+follows. Choose~$y$ such that~$\neg P(y)$, if such exists; otherwise
+$\all{x}P(x)$ is true. Either way the theorem holds.
+
+The formal proof does not conform in any obvious way to the sketch given
+above. The key inference is the first one, \ttindex{exCI}; this classical
+version of~$(\exists I)$ allows multiple instantiation of the quantifier.
+\begin{ttbox}
+goal FOL.thy "EX y. ALL x. P(y)-->P(x)";
+{\out Level 0}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. EX y. ALL x. P(y) --> P(x)}
+\ttbreak
+by (resolve_tac [exCI] 1);
+{\out Level 1}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. ALL y. ~ (ALL x. P(y) --> P(x)) ==> ALL x. P(?a) --> P(x)}
+\end{ttbox}
+We now can either exhibit a term {\tt?a} to satisfy the conclusion of
+subgoal~1, or produce a contradiction from the assumption. The next
+steps routinely break down the conclusion and assumption.
+\begin{ttbox}
+by (resolve_tac [allI] 1);
+{\out Level 2}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. ALL y. ~ (ALL x. P(y) --> P(x)) ==> P(?a) --> P(x)}
+\ttbreak
+by (resolve_tac [impI] 1);
+{\out Level 3}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. [| ALL y. ~ (ALL x. P(y) --> P(x)); P(?a) |] ==> P(x)}
+\ttbreak
+by (eresolve_tac [allE] 1);
+{\out Level 4}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. [| P(?a); ~ (ALL xa. P(?y3(x)) --> P(xa)) |] ==> P(x)}
+\end{ttbox}
+In classical logic, a negated assumption is equivalent to a conclusion. To
+get this effect, we create a swapped version of $(\forall I)$ and apply it
+using \ttindex{eresolve_tac}; we could equivalently have applied~$(\forall
+I)$ using \ttindex{swap_res_tac}.
+\begin{ttbox}
+allI RSN (2,swap);
+{\out val it = "[| ~ (ALL x. ?P1(x)); !!x. ~ ?Q ==> ?P1(x) |] ==> ?Q" : thm}
+by (eresolve_tac [it] 1);
+{\out Level 5}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x xa. [| P(?a); ~ P(x) |] ==> P(?y3(x)) --> P(xa)}
+\end{ttbox}
+The previous conclusion, {\tt P(x)}, has become a negated assumption;
+we apply~$({\imp}I)$:
+\begin{ttbox}
+by (resolve_tac [impI] 1);
+{\out Level 6}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x xa. [| P(?a); ~ P(x); P(?y3(x)) |] ==> P(xa)}
+\end{ttbox}
+The subgoal has three assumptions. We produce a contradiction between the
+assumptions~\verb|~P(x)| and~{\tt P(?y3(x))}. The proof never instantiates
+the unknown~{\tt?a}.
+\begin{ttbox}
+by (eresolve_tac [notE] 1);
+{\out Level 7}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x xa. [| P(?a); P(?y3(x)) |] ==> P(x)}
+\ttbreak
+by (assume_tac 1);
+{\out Level 8}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out No subgoals!}
+\end{ttbox}
+The civilized way to prove this theorem is through \ttindex{best_tac},
+supplying the classical version of~$(\exists I)$:
+\begin{ttbox}
+goal FOL.thy "EX y. ALL x. P(y)-->P(x)";
+{\out Level 0}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out 1. EX y. ALL x. P(y) --> P(x)}
+by (best_tac FOL_dup_cs 1);
+{\out Level 1}
+{\out EX y. ALL x. P(y) --> P(x)}
+{\out No subgoals!}
+\end{ttbox}
+If this theorem seems counterintuitive, then perhaps you are an
+intuitionist. In constructive logic, proving $\ex{y}\all{x}P(y)\imp P(x)$
+requires exhibiting a particular term~$t$ such that $\all{x}P(t)\imp P(x)$,
+which we cannot do without further knowledge about~$P$.
+
+
+\section{Derived rules and the classical tactics}
+Classical first-order logic can be extended with the propositional
+connective $if(P,Q,R)$, where
+$$ if(P,Q,R) \equiv P\conj Q \disj \neg P \conj R. \eqno(if) $$
+Theorems about $if$ can be proved by treating this as an abbreviation,
+replacing $if(P,Q,R)$ by $P\conj Q \disj \neg P \conj R$ in subgoals. But
+this duplicates~$P$, causing an exponential blowup and an unreadable
+formula. Introducing further abbreviations makes the problem worse.
+
+Natural deduction demands rules that introduce and eliminate $if(P,Q,R)$
+directly, without reference to its definition. The simple identity
+\[ if(P,Q,R) \bimp (P\imp Q)\conj (\neg P\imp R) \]
+suggests that the
+$if$-introduction rule should be
+\[ \infer[({if}\,I)]{if(P,Q,R)}{\infer*{Q}{P} & \infer*{R}{\neg P}} \]
+The $if$-elimination rule reflects the definition of $if(P,Q,R)$ and the
+elimination rules for~$\disj$ and~$\conj$.
+\[ \infer[({if}\,E)]{S}{if(P,Q,R) & \infer*{S}{[P,Q]}
+ & \infer*{S}{[\neg P,R]}}
+\]
+Having made these plans, we get down to work with Isabelle. The theory of
+classical logic, \ttindex{FOL}, is extended with the constant
+$if::[o,o,o]\To o$. The axiom \ttindexbold{if_def} asserts the
+equation~$(if)$.
+\begin{ttbox}
+If = FOL +
+consts if :: "[o,o,o]=>o"
+rules if_def "if(P,Q,R) == P&Q | ~P&R"
+end
+\end{ttbox}
+The derivations of the introduction and elimination rules demonstrate the
+methods for rewriting with definitions. Classical reasoning is required,
+so we use \ttindex{fast_tac}.
+
+
+\subsection{Deriving the introduction rule}
+The introduction rule, given the premises $P\Imp Q$ and $\neg P\Imp R$,
+concludes $if(P,Q,R)$. We propose the conclusion as the main goal
+using~\ttindex{goalw}, which uses {\tt if_def} to rewrite occurrences
+of $if$ in the subgoal.
+\begin{ttbox}
+val prems = goalw If.thy [if_def]
+ "[| P ==> Q; ~ P ==> R |] ==> if(P,Q,R)";
+{\out Level 0}
+{\out if(P,Q,R)}
+{\out 1. P & Q | ~ P & R}
+\end{ttbox}
+The premises (bound to the {\ML} variable {\tt prems}) are passed as
+introduction rules to \ttindex{fast_tac}:
+\begin{ttbox}
+by (fast_tac (FOL_cs addIs prems) 1);
+{\out Level 1}
+{\out if(P,Q,R)}
+{\out No subgoals!}
+val ifI = result();
+\end{ttbox}
+
+
+\subsection{Deriving the elimination rule}
+The elimination rule has three premises, two of which are themselves rules.
+The conclusion is simply $S$.
+\begin{ttbox}
+val major::prems = goalw If.thy [if_def]
+ "[| if(P,Q,R); [| P; Q |] ==> S; [| ~ P; R |] ==> S |] ==> S";
+{\out Level 0}
+{\out S}
+{\out 1. S}
+\end{ttbox}
+The major premise contains an occurrence of~$if$, but the version returned
+by \ttindex{goalw} (and bound to the {\ML} variable~{\tt major}) has the
+definition expanded. Now \ttindex{cut_facts_tac} inserts~{\tt major} as an
+assumption in the subgoal, so that \ttindex{fast_tac} can break it down.
+\begin{ttbox}
+by (cut_facts_tac [major] 1);
+{\out Level 1}
+{\out S}
+{\out 1. P & Q | ~ P & R ==> S}
+\ttbreak
+by (fast_tac (FOL_cs addIs prems) 1);
+{\out Level 2}
+{\out S}
+{\out No subgoals!}
+val ifE = result();
+\end{ttbox}
+As you may recall from {\em Introduction to Isabelle}, there are other
+ways of treating definitions when deriving a rule. We can start the
+proof using \ttindex{goal}, which does not expand definitions, instead of
+\ttindex{goalw}. We can use \ttindex{rewrite_goals_tac}
+to expand definitions in the subgoals --- perhaps after calling
+\ttindex{cut_facts_tac} to insert the rule's premises. We can use
+\ttindex{rewrite_rule}, which is a meta-inference rule, to expand
+definitions in the premises directly.
+
+
+\subsection{Using the derived rules}
+The rules just derived have been saved with the {\ML} names \ttindex{ifI}
+and~\ttindex{ifE}. They permit natural proofs of theorems such as the
+following:
+\begin{eqnarray*}
+ if(P, if(Q,A,B), if(Q,C,D)) & \bimp & if(Q,if(P,A,C),if(P,B,D)) \\
+ if(if(P,Q,R), A, B) & \bimp & if(P,if(Q,A,B),if(R,A,B))
+\end{eqnarray*}
+Proofs also require the classical reasoning rules and the $\bimp$
+introduction rule (called~\ttindex{iffI}: do not confuse with~\ttindex{ifI}).
+
+To display the $if$-rules in action, let us analyse a proof step by step.
+\begin{ttbox}
+goal If.thy
+ "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))";
+{\out Level 0}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+\ttbreak
+by (resolve_tac [iffI] 1);
+{\out Level 1}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. if(P,if(Q,A,B),if(Q,C,D)) ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 2. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
+\end{ttbox}
+The $if$-elimination rule can be applied twice in succession.
+\begin{ttbox}
+by (eresolve_tac [ifE] 1);
+{\out Level 2}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. [| P; if(Q,A,B) |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 2. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 3. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
+\ttbreak
+by (eresolve_tac [ifE] 1);
+{\out Level 3}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. [| P; Q; A |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 2. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 3. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 4. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
+\end{ttbox}
+
+In the first two subgoals, all formulae have been reduced to atoms. Now
+$if$-introduction can be applied. Observe how the $if$-rules break down
+occurrences of $if$ when they become the outermost connective.
+\begin{ttbox}
+by (resolve_tac [ifI] 1);
+{\out Level 4}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. [| P; Q; A; Q |] ==> if(P,A,C)}
+{\out 2. [| P; Q; A; ~ Q |] ==> if(P,B,D)}
+{\out 3. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 4. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 5. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
+\ttbreak
+by (resolve_tac [ifI] 1);
+{\out Level 5}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. [| P; Q; A; Q; P |] ==> A}
+{\out 2. [| P; Q; A; Q; ~ P |] ==> C}
+{\out 3. [| P; Q; A; ~ Q |] ==> if(P,B,D)}
+{\out 4. [| P; ~ Q; B |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 5. [| ~ P; if(Q,C,D) |] ==> if(Q,if(P,A,C),if(P,B,D))}
+{\out 6. if(Q,if(P,A,C),if(P,B,D)) ==> if(P,if(Q,A,B),if(Q,C,D))}
+\end{ttbox}
+Where do we stand? The first subgoal holds by assumption; the second and
+third, by contradiction. This is getting tedious. Let us revert to the
+initial proof state and let \ttindex{fast_tac} solve it. The classical
+rule set {\tt if_cs} contains the rules of~{\FOL} plus the derived rules
+for~$if$.
+\begin{ttbox}
+choplev 0;
+{\out Level 0}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out 1. if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+val if_cs = FOL_cs addSIs [ifI] addSEs[ifE];
+by (fast_tac if_cs 1);
+{\out Level 1}
+{\out if(P,if(Q,A,B),if(Q,C,D)) <-> if(Q,if(P,A,C),if(P,B,D))}
+{\out No subgoals!}
+\end{ttbox}
+This tactic also solves the other example.
+\begin{ttbox}
+goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))";
+{\out Level 0}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+\ttbreak
+by (fast_tac if_cs 1);
+{\out Level 1}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\subsection{Derived rules versus definitions}
+Dispensing with the derived rules, we can treat $if$ as an
+abbreviation, and let \ttindex{fast_tac} prove the expanded formula. Let
+us redo the previous proof:
+\begin{ttbox}
+choplev 0;
+{\out Level 0}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+\ttbreak
+by (rewrite_goals_tac [if_def]);
+{\out Level 1}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+{\out 1. (P & Q | ~ P & R) & A | ~ (P & Q | ~ P & R) & B <->}
+{\out P & (Q & A | ~ Q & B) | ~ P & (R & A | ~ R & B)}
+\ttbreak
+by (fast_tac FOL_cs 1);
+{\out Level 2}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,A,B))}
+{\out No subgoals!}
+\end{ttbox}
+Expanding definitions reduces the extended logic to the base logic. This
+approach has its merits --- especially if the prover for the base logic is
+good --- but can be slow. In these examples, proofs using {\tt if_cs} (the
+derived rules) run about six times faster than proofs using {\tt FOL_cs}.
+
+Expanding definitions also complicates error diagnosis. Suppose we are having
+difficulties in proving some goal. If by expanding definitions we have
+made it unreadable, then we have little hope of diagnosing the problem.
+
+Attempts at program verification often yield invalid assertions.
+Let us try to prove one:
+\begin{ttbox}
+goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))";
+{\out Level 0}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+by (fast_tac FOL_cs 1);
+{\out by: tactic failed}
+\end{ttbox}
+This failure message is uninformative, but we can get a closer look at the
+situation by applying \ttindex{step_tac}.
+\begin{ttbox}
+by (REPEAT (step_tac if_cs 1));
+{\out Level 1}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+{\out 1. [| A; ~ P; R; ~ P; R |] ==> B}
+{\out 2. [| B; ~ P; ~ R; ~ P; ~ R |] ==> A}
+{\out 3. [| ~ P; R; B; ~ P; R |] ==> A}
+{\out 4. [| ~ P; ~ R; A; ~ B; ~ P |] ==> R}
+\end{ttbox}
+Subgoal~1 is unprovable and yields a countermodel: $P$ and~$B$ are false
+while~$R$ and~$A$ are true. This truth assignment reduces the main goal to
+$true\bimp false$, which is of course invalid.
+
+We can repeat this analysis by expanding definitions, using just
+the rules of {\FOL}:
+\begin{ttbox}
+choplev 0;
+{\out Level 0}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+{\out 1. if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+\ttbreak
+by (rewrite_goals_tac [if_def]);
+{\out Level 1}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+{\out 1. (P & Q | ~ P & R) & A | ~ (P & Q | ~ P & R) & B <->}
+{\out P & (Q & A | ~ Q & B) | ~ P & (R & B | ~ R & A)}
+by (fast_tac FOL_cs 1);
+{\out by: tactic failed}
+\end{ttbox}
+Again we apply \ttindex{step_tac}:
+\begin{ttbox}
+by (REPEAT (step_tac FOL_cs 1));
+{\out Level 2}
+{\out if(if(P,Q,R),A,B) <-> if(P,if(Q,A,B),if(R,B,A))}
+{\out 1. [| A; ~ P; R; ~ P; R; ~ False |] ==> B}
+{\out 2. [| A; ~ P; R; R; ~ False; ~ B; ~ B |] ==> Q}
+{\out 3. [| B; ~ P; ~ R; ~ P; ~ A |] ==> R}
+{\out 4. [| B; ~ P; ~ R; ~ Q; ~ A |] ==> R}
+{\out 5. [| B; ~ R; ~ P; ~ A; ~ R; Q; ~ False |] ==> A}
+{\out 6. [| ~ P; R; B; ~ P; R; ~ False |] ==> A}
+{\out 7. [| ~ P; ~ R; A; ~ B; ~ R |] ==> P}
+{\out 8. [| ~ P; ~ R; A; ~ B; ~ R |] ==> Q}
+\end{ttbox}
+Subgoal~1 yields the same countermodel as before. But each proof step has
+taken six times as long, and the final result contains twice as many subgoals.
+
+Expanding definitions causes a great increase in complexity. This is why
+the classical prover has been designed to accept derived rules.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/HOL-eg.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,151 @@
+(**** HOL examples -- process using Doc/tout HOL-eg.txt ****)
+
+Pretty.setmargin 72; (*existing macros just allow this margin*)
+print_depth 0;
+
+
+(*** Conjunction rules ***)
+
+val prems = goal HOL_Rule.thy "[| P; Q |] ==> P&Q";
+by (resolve_tac [and_def RS ssubst] 1);
+by (resolve_tac [allI] 1);
+by (resolve_tac [impI] 1);
+by (eresolve_tac [mp RS mp] 1);
+by (REPEAT (resolve_tac prems 1));
+val conjI = result();
+
+val prems = goal HOL_Rule.thy "[| P & Q |] ==> P";
+prths (prems RL [and_def RS subst]);
+prths (prems RL [and_def RS subst] RL [spec] RL [mp]);
+by (resolve_tac it 1);
+by (REPEAT (ares_tac [impI] 1));
+val conjunct1 = result();
+
+
+(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***)
+
+goal Set.thy "~ ?S : range(f :: 'a=>'a set)";
+by (resolve_tac [notI] 1);
+by (eresolve_tac [rangeE] 1);
+by (eresolve_tac [equalityCE] 1);
+by (dresolve_tac [CollectD] 1);
+by (contr_tac 1);
+by (swap_res_tac [CollectI] 1);
+by (assume_tac 1);
+
+choplev 0;
+by (best_tac (set_cs addSEs [equalityCE]) 1);
+
+
+goal Set.thy "! f:: 'a=>'a set. ! x. ~ f(x) = ?S(f)";
+by (REPEAT (resolve_tac [allI,notI] 1));
+by (eresolve_tac [equalityCE] 1);
+by (dresolve_tac [CollectD] 1);
+by (contr_tac 1);
+by (swap_res_tac [CollectI] 1);
+by (assume_tac 1);
+
+choplev 0;
+by (best_tac (set_cs addSEs [equalityCE]) 1);
+
+
+goal Set.thy "~ (? f:: 'a=>'a set. ! S. ? a. f(a) = S)";
+by (best_tac (set_cs addSEs [equalityCE]) 1);
+
+
+
+
+> val prems = goal HOL_Rule.thy "[| P; Q |] ==> P&Q";
+Level 0
+P & Q
+ 1. P & Q
+> by (resolve_tac [and_def RS ssubst] 1);
+Level 1
+P & Q
+ 1. ! R. (P --> Q --> R) --> R
+> by (resolve_tac [allI] 1);
+Level 2
+P & Q
+ 1. !!R. (P --> Q --> R) --> R
+> by (resolve_tac [impI] 1);
+Level 3
+P & Q
+ 1. !!R. P --> Q --> R ==> R
+> by (eresolve_tac [mp RS mp] 1);
+Level 4
+P & Q
+ 1. !!R. P
+ 2. !!R. Q
+> by (REPEAT (resolve_tac prems 1));
+Level 5
+P & Q
+No subgoals!
+
+
+
+> val prems = goal HOL_Rule.thy "[| P & Q |] ==> P";
+Level 0
+P
+ 1. P
+> prths (prems RL [and_def RS subst]);
+! R. (P --> Q --> R) --> R [P & Q]
+P & Q [P & Q]
+
+> prths (prems RL [and_def RS subst] RL [spec] RL [mp]);
+P --> Q --> ?Q ==> ?Q [P & Q]
+
+> by (resolve_tac it 1);
+Level 1
+P
+ 1. P --> Q --> P
+> by (REPEAT (ares_tac [impI] 1));
+Level 2
+P
+No subgoals!
+
+
+
+
+> goal Set.thy "~ ?S : range(f :: 'a=>'a set)";
+Level 0
+~?S : range(f)
+ 1. ~?S : range(f)
+> by (resolve_tac [notI] 1);
+Level 1
+~?S : range(f)
+ 1. ?S : range(f) ==> False
+> by (eresolve_tac [rangeE] 1);
+Level 2
+~?S : range(f)
+ 1. !!x. ?S = f(x) ==> False
+> by (eresolve_tac [equalityCE] 1);
+Level 3
+~?S : range(f)
+ 1. !!x. [| ?c3(x) : ?S; ?c3(x) : f(x) |] ==> False
+ 2. !!x. [| ~?c3(x) : ?S; ~?c3(x) : f(x) |] ==> False
+> by (dresolve_tac [CollectD] 1);
+Level 4
+~{x. ?P7(x)} : range(f)
+ 1. !!x. [| ?c3(x) : f(x); ?P7(?c3(x)) |] ==> False
+ 2. !!x. [| ~?c3(x) : {x. ?P7(x)}; ~?c3(x) : f(x) |] ==> False
+> by (contr_tac 1);
+Level 5
+~{x. ~x : f(x)} : range(f)
+ 1. !!x. [| ~x : {x. ~x : f(x)}; ~x : f(x) |] ==> False
+> by (swap_res_tac [CollectI] 1);
+Level 6
+~{x. ~x : f(x)} : range(f)
+ 1. !!x. [| ~x : f(x); ~False |] ==> ~x : f(x)
+> by (assume_tac 1);
+Level 7
+~{x. ~x : f(x)} : range(f)
+No subgoals!
+
+> choplev 0;
+Level 0
+~?S : range(f)
+ 1. ~?S : range(f)
+> by (best_tac (set_cs addSEs [equalityCE]) 1);
+Level 1
+~{x. ~x : f(x)} : range(f)
+No subgoals!
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/HOL-rules.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,403 @@
+ruleshell.ML lemmas.ML set.ML fun.ML subset.ML equalities.ML prod.ML sum.ML wf.ML mono.ML fixedpt.ML nat.ML list.ML
+----------------------------------------------------------------
+ruleshell.ML
+
+\idx{refl} t = t::'a
+\idx{subst} [| s = t; P(s) |] ==> P(t::'a)
+\idx{abs},!!x::'a. f(x)::'b = g(x)) ==> (%x.f(x)) = (%x.g(x)))
+\idx{disch} (P ==> Q) ==> P-->Q
+\idx{mp} [| P-->Q; P |] ==> Q
+
+\idx{True_def} True = ((%x.x)=(%x.x))
+\idx{All_def} All = (%P. P = (%x.True))
+\idx{Ex_def} Ex = (%P. P(Eps(P)))
+\idx{False_def} False = (!P.P)
+\idx{not_def} not = (%P. P-->False)
+\idx{and_def} op & = (%P Q. !R. (P-->Q-->R) --> R)
+\idx{or_def} op | = (%P Q. !R. (P-->R) --> (Q-->R) --> R)
+\idx{Ex1_def} Ex1 == (%P. ? x. P(x) & (! y. P(y) --> y=x))
+
+\idx{iff} (P-->Q) --> (Q-->P) --> (P=Q)
+\idx{True_or_False} (P=True) | (P=False)
+\idx{select} P(x::'a) --> P(Eps(P))
+
+\idx{Inv_def} Inv = (%(f::'a=>'b) y. @x. f(x)=y)
+\idx{o_def} op o = (%(f::'b=>'c) g (x::'a). f(g(x)))
+\idx{Cond_def} Cond = (%P x y.@z::'a. (P=True --> z=x) & (P=False --> z=y))
+
+----------------------------------------------------------------
+lemmas.ML
+
+\idx{sym} s=t ==> t=s
+\idx{trans} [| r=s; s=t |] ==> r=t
+\idx{box_equals}
+ [| a=b; a=c; b=d |] ==> c=d
+\idx{ap_term} s=t ==> f(s)=f(t)
+\idx{ap_thm} s::'a=>'b = t ==> s(x)=t(x)
+\idx{cong}
+ [| f = g; x::'a = y |] ==> f(x) = g(y)
+\idx{iffI}
+ [| P ==> Q; Q ==> P |] ==> P=Q
+\idx{iffD1} [| P=Q; Q |] ==> P
+\idx{iffE}
+ [| P=Q; [| P --> Q; Q --> P |] ==> R |] ==> R
+\idx{eqTrueI} P ==> P=True
+\idx{eqTrueE} P=True ==> P
+\idx{allI} (!!x::'a. P(x)) ==> !x. P(x)
+\idx{spec} !x::'a.P(x) ==> P(x)
+\idx{allE} [| !x.P(x); P(x) ==> R |] ==> R
+\idx{all_dupE}
+ [| ! x.P(x); [| P(x); ! x.P(x) |] ==> R
+ |] ==> R
+\idx{FalseE} False ==> P
+\idx{False_neq_True} False=True ==> P
+\idx{notI} (P ==> False) ==> ~P
+\idx{notE} [| ~P; P |] ==> R
+\idx{impE} [| P-->Q; P; Q ==> R |] ==> R
+\idx{rev_mp} [| P; P --> Q |] ==> Q
+\idx{contrapos} [| ~Q; P==>Q |] ==> ~P
+\idx{exI} P(x) ==> ? x::'a.P(x)
+\idx{exE} [| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q
+
+\idx{conjI} [| P; Q |] ==> P&Q
+\idx{conjunct1} [| P & Q |] ==> P
+\idx{conjunct2} [| P & Q |] ==> Q
+\idx{conjE} [| P&Q; [| P; Q |] ==> R |] ==> R
+\idx{disjI1} P ==> P|Q
+\idx{disjI2} Q ==> P|Q
+\idx{disjE} [| P | Q; P ==> R; Q ==> R |] ==> R
+\idx{ccontr} (~P ==> False) ==> P
+\idx{classical} (~P ==> P) ==> P
+\idx{notnotD} ~~P ==> P
+\idx{ex1I}
+ [| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)
+\idx{ex1E}
+ [| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R |] ==> R
+\idx{select_equality}
+ [| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a
+\idx{disjCI} (~Q ==> P) ==> P|Q
+\idx{excluded_middle} ~P | P
+\idx{impCE} [| P-->Q; ~P ==> R; Q ==> R |] ==> R
+\idx{iffCE}
+ [| P=Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R
+\idx{exCI} (! x. ~P(x) ==> P(a)) ==> ? x.P(x)
+\idx{swap} ~P ==> (~Q ==> P) ==> Q
+
+----------------------------------------------------------------
+simpdata.ML
+
+\idx{if_True} Cond(True,x,y) = x
+\idx{if_False} Cond(False,x,y) = y
+\idx{if_P} P ==> Cond(P,x,y) = x
+\idx{if_not_P} ~P ==> Cond(P,x,y) = y
+\idx{expand_if}
+ P(Cond(Q,x,y)) = ((Q --> P(x)) & (~Q --> P(y)))
+
+----------------------------------------------------------------
+\idx{set.ML}
+
+\idx{CollectI} [| P(a) |] ==> a : \{x.P(x)\}
+\idx{CollectD} [| a : \{x.P(x)\} |] ==> P(a)
+\idx{set_ext} [| !!x. (x:A) = (x:B) |] ==> A = B
+
+\idx{Ball_def} Ball(A,P) == ! x. x:A --> P(x)
+\idx{Bex_def} Bex(A,P) == ? x. x:A & P(x)
+\idx{subset_def} A <= B == ! x:A. x:B
+\idx{Un_def} A Un B == \{x.x:A | x:B\}
+\idx{Int_def} A Int B == \{x.x:A & x:B\}
+\idx{Compl_def} Compl(A) == \{x. ~x:A\}
+\idx{Inter_def} Inter(S) == \{x. ! A:S. x:A\}
+\idx{Union_def} Union(S) == \{x. ? A:S. x:A\}
+\idx{INTER_def} INTER(A,B) == \{y. ! x:A. y: B(x)\}
+\idx{UNION_def} UNION(A,B) == \{y. ? x:A. y: B(x)\}
+\idx{mono_def} mono(f) == (!A B. A <= B --> f(A) <= f(B))
+\idx{image_def} f``A == \{y. ? x:A. y=f(x)\}
+\idx{singleton_def} \{a\} == \{x.x=a\}
+\idx{range_def} range(f) == \{y. ? x. y=f(x)\}
+\idx{One_One_def} One_One(f) == ! x y. f(x)=f(y) --> x=y
+\idx{One_One_on_def} One_One_on(f,A) == !x y. x:A --> y:A --> f(x)=f(y) --> x=y
+\idx{Onto_def} Onto(f) == ! y. ? x. y=f(x)
+
+
+\idx{Collect_cong} [| !!x. P(x)=Q(x) |] ==> \{x. P(x)\} = \{x. Q(x)\}
+
+\idx{ballI} [| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)
+\idx{bspec} [| ! x:A. P(x); x:A |] ==> P(x)
+\idx{ballE} [| ! x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q
+
+\idx{bexI} [| P(x); x:A |] ==> ? x:A. P(x)
+\idx{bexCI} [| ! x:A. ~P(x) ==> P(a); a:A |] ==> ? x:A.P(x)
+\idx{bexE} [| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q
+
+\idx{ball_cong}
+ [| A=A'; !!x. x:A' ==> P(x) = P'(x) |] ==>
+ (! x:A. P(x)) = (! x:A'. P'(x))
+
+\idx{bex_cong}
+ [| A=A'; !!x. x:A' ==> P(x) = P'(x) |] ==>
+ (? x:A. P(x)) = (? x:A'. P'(x))
+
+\idx{subsetI} (!!x.x:A ==> x:B) ==> A <= B
+\idx{subsetD} [| A <= B; c:A |] ==> c:B
+\idx{subsetCE} [| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P
+
+\idx{subset_refl} A <= A
+\idx{subset_antisym} [| A <= B; B <= A |] ==> A = B
+\idx{subset_trans} [| A<=B; B<=C |] ==> A<=C
+
+\idx{equalityD1} A = B ==> A<=B
+\idx{equalityD2} A = B ==> B<=A
+\idx{equalityE} [| A = B; [| A<=B; B<=A |] ==> P |] ==> P
+
+\idx{singletonI} a : \{a\}
+\idx{singletonD} b : \{a\} ==> b=a
+
+\idx{imageI} [| x:A |] ==> f(x) : f``A
+\idx{imageE} [| b : f``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P
+
+\idx{rangeI} f(x) : range(f)
+\idx{rangeE} [| b : range(f); !!x.[| b=f(x) |] ==> P |] ==> P
+
+\idx{UnionI} [| X:C; A:X |] ==> A : Union(C)
+\idx{UnionE} [| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R
+
+\idx{InterI} [| !!X. X:C ==> A:X |] ==> A : Inter(C)
+\idx{InterD} [| A : Inter(C); X:C |] ==> A:X
+\idx{InterE} [| A : Inter(C); A:X ==> R; ~ X:C ==> R |] ==> R
+
+\idx{UN_I} [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))
+\idx{UN_E} [| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R
+
+\idx{INT_I} (!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))
+\idx{INT_D} [| b : (INT x:A. B(x)); a:A |] ==> b: B(a)
+\idx{INT_E} [| b : (INT x:A. B(x)); b: B(a) ==> R; ~ a:A ==> R |] ==> R
+
+\idx{UnI1} c:A ==> c : A Un B
+\idx{UnI2} c:B ==> c : A Un B
+\idx{UnCI} (~c:B ==> c:A) ==> c : A Un B
+\idx{UnE} [| c : A Un B; c:A ==> P; c:B ==> P |] ==> P
+
+\idx{IntI} [| c:A; c:B |] ==> c : A Int B
+\idx{IntD1} c : A Int B ==> c:A
+\idx{IntD2} c : A Int B ==> c:B
+\idx{IntE} [| c : A Int B; [| c:A; c:B |] ==> P |] ==> P
+
+\idx{ComplI} [| c:A ==> False |] ==> c : Compl(A)
+\idx{ComplD} [| c : Compl(A) |] ==> ~c:A
+
+\idx{monoI} [| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)
+\idx{monoD} [| mono(f); A <= B |] ==> f(A) <= f(B)
+
+
+----------------------------------------------------------------
+\idx{fun.ML}
+
+\idx{One_OneI} [| !! x y. f(x) = f(y) ==> x=y |] ==> One_One(f)
+\idx{One_One_inverseI} (!!x. g(f(x)) = x) ==> One_One(f)
+\idx{One_OneD} [| One_One(f); f(x) = f(y) |] ==> x=y
+
+\idx{Inv_f_f} One_One(f) ==> Inv(f,f(x)) = x
+\idx{f_Inv_f} y : range(f) ==> f(Inv(f,y)) = y
+
+\idx{Inv_injective}
+ [| Inv(f,x)=Inv(f,y); x: range(f); y: range(f) |] ==> x=y
+
+\idx{One_One_onI}
+ (!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> One_One_on(f,A)
+
+\idx{One_One_on_inverseI}
+ (!!x. x:A ==> g(f(x)) = x) ==> One_One_on(f,A)
+
+\idx{One_One_onD}
+ [| One_One_on(f,A); f(x)=f(y); x:A; y:A |] ==> x=y
+
+\idx{One_One_on_contraD}
+ [| One_One_on(f,A); ~x=y; x:A; y:A |] ==> ~ f(x)=f(y)
+
+
+----------------------------------------------------------------
+\idx{subset.ML}
+
+\idx{Union_upper} B:A ==> B <= Union(A)
+\idx{Union_least} [| !!X. X:A ==> X<=C |] ==> Union(A) <= C
+
+\idx{Inter_lower} B:A ==> Inter(A) <= B
+\idx{Inter_greatest} [| !!X. X:A ==> C<=X |] ==> C <= Inter(A)
+
+\idx{Un_upper1} A <= A Un B
+\idx{Un_upper2} B <= A Un B
+\idx{Un_least} [| A<=C; B<=C |] ==> A Un B <= C
+
+\idx{Int_lower1} A Int B <= A
+\idx{Int_lower2} A Int B <= B
+\idx{Int_greatest} [| C<=A; C<=B |] ==> C <= A Int B
+
+
+----------------------------------------------------------------
+\idx{equalities.ML}
+
+\idx{Int_absorb} A Int A = A
+\idx{Int_commute} A Int B = B Int A
+\idx{Int_assoc} (A Int B) Int C = A Int (B Int C)
+\idx{Int_Un_distrib} (A Un B) Int C = (A Int C) Un (B Int C)
+
+\idx{Un_absorb} A Un A = A
+\idx{Un_commute} A Un B = B Un A
+\idx{Un_assoc} (A Un B) Un C = A Un (B Un C)
+\idx{Un_Int_distrib} (A Int B) Un C = (A Un C) Int (B Un C)
+
+\idx{Compl_disjoint} A Int Compl(A) = \{x.False\}
+\idx{Compl_partition A Un Compl(A) = \{x.True\}
+\idx{double_complement} Compl(Compl(A)) = A
+
+
+\idx{Compl_Un} Compl(A Un B) = Compl(A) Int Compl(B)
+\idx{Compl_Int} Compl(A Int B) = Compl(A) Un Compl(B)
+
+\idx{Union_Un_distrib} Union(A Un B) = Union(A) Un Union(B)
+\idx{Int_Union_image} A Int Union(B) = (UN C:B. A Int C)
+\idx{Un_Union_image} (UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)
+
+\idx{Inter_Un_distrib} Inter(A Un B) = Inter(A) Int Inter(B)
+\idx{Un_Inter_image} A Un Inter(B) = (INT C:B. A Un C)
+\idx{Int_Inter_image} (INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)
+
+
+----------------------------------------------------------------
+prod.ML
+
+ mixfix = [ Delimfix((1<_,/_>), ['a,'b] => ('a,'b)prod, Pair),
+ TInfixl(*, prod, 20) ],
+thy = extend_theory Set.thy Prod
+ [([prod],([[term],[term]],term))],
+ ([fst], 'a * 'b => 'a),
+ ([snd], 'a * 'b => 'b),
+ ([split], ['a * 'b, ['a,'b]=>'c] => 'c)],
+\idx{fst_def} fst(p) == @a. ? b. p = <a,b>),
+\idx{snd_def} snd(p) == @b. ? a. p = <a,b>),
+\idx{split_def} split(p,c) == c(fst(p),snd(p)))
+
+\idx{Pair_inject} [| <a, b> = <a',b'>; [| a=a'; b=b' |] ==> R |] ==> R
+
+\idx{fst_conv} fst(<a,b>) = a
+\idx{snd_conv} snd(<a,b>) = b
+\idx{split_conv} split(<a,b>, c) = c(a,b)
+
+\idx{surjective_pairing} p = <fst(p),snd(p)>
+
+----------------------------------------------------------------
+sum.ML
+
+ mixfix = [TInfixl(+, sum, 10)],
+thy = extend_theory Prod.thy sum
+ [([sum], ([[term],[term]],term))],
+ [Inl], 'a => 'a+'b),
+ [Inr], 'b => 'a+'b),
+ [when], ['a+'b, 'a=>'c, 'b=>'c] =>'c)],
+\idx{when_def} when == (%p f g. @z. (!x. p=Inl(x) --> z=f(x))
+ & (!y. p=Inr(y) --> z=g(y))))
+
+\idx{Inl_not_Inr} ~ (Inl(a) = Inr(b))
+
+\idx{One_One_Inl} One_One(Inl)
+
+\idx{One_One_Inr} One_One(Inr)
+
+\idx{when_Inl_conv} when(Inl(x), f, g) = f(x)
+
+\idx{when_Inr_conv} when(Inr(x), f, g) = g(x)
+
+\idx{sumE}
+ [| !!x::'a. P(Inl(x)); !!y::'b. P(Inr(y))
+ |] ==> P(s)
+
+\idx{surjective_sum} when(s, %x::'a. f(Inl(x)), %y::'b. f(Inr(y))) = f(s)
+
+
+????????????????????????????????????????????????????????????????
+trancl?
+
+----------------------------------------------------------------
+nat.ML
+
+ Sext\{mixfix=[Delimfix(0, nat, 0),
+ Infixl(<,[nat,nat] => bool,50)],
+thy = extend_theory Trancl.thy Nat
+[nat], ([],term))
+[nat_case], [nat, 'a, nat=>'a] =>'a),
+[pred_nat],nat*nat) set),
+[nat_rec], [nat, 'a, [nat, 'a]=>'a] => 'a)
+
+\idx{nat_case_def} nat_case == (%n a f. @z. (n=0 --> z=a)
+ & (!x. n=Suc(x) --> z=f(x)))),
+\idx{pred_nat_def} pred_nat == \{p. ? n. p = <n, Suc(n)>\} ),
+\idx{less_def} m<n == <m,n>:trancl(pred_nat)),
+\idx{nat_rec_def}
+ nat_rec(n,c,d) == wfrec(trancl(pred_nat),
+ %rec l. nat_case(l, c, %m. d(m,rec(m))),
+ n) )
+
+\idx{nat_induct} [| P(0); !!k. [| P(k) |] ==> P(Suc(k)) |] ==> P(n)
+
+
+\idx{Suc_not_Zero} ~ (Suc(m) = 0)
+\idx{One_One_Suc} One_One(Suc)
+\idx{n_not_Suc_n} ~(n=Suc(n))
+
+\idx{nat_case_0_conv} nat_case(0, a, f) = a
+
+\idx{nat_case_Suc_conv} nat_case(Suc(k), a, f) = f(k)
+
+\idx{pred_natI} <n, Suc(n)> : pred_nat
+\idx{pred_natE}
+ [| p : pred_nat; !!x n. [| p = <n, Suc(n)> |] ==> R
+ |] ==> R
+
+\idx{wf_pred_nat} wf(pred_nat)
+
+\idx{nat_rec_0_conv} nat_rec(0,c,h) = c
+
+\idx{nat_rec_Suc_conv} nat_rec(Suc(n), c, h) = h(n, nat_rec(n,c,h))
+
+
+(*** Basic properties of less than ***)
+\idx{less_trans} [| i<j; j<k |] ==> i<k
+\idx{lessI} n < Suc(n)
+\idx{zero_less_Suc} 0 < Suc(n)
+
+\idx{less_not_sym} n<m --> ~m<n
+\idx{less_not_refl} ~ (n<n)
+\idx{not_less0} ~ (n<0)
+
+\idx{Suc_less_eq} (Suc(m) < Suc(n)) = (m<n)
+\idx{less_induct} [| !!n. [| ! m. m<n --> P(m) |] ==> P(n) |] ==> P(n)
+
+\idx{less_linear} m<n | m=n | n<m
+
+
+----------------------------------------------------------------
+list.ML
+
+ [([list], ([[term]],term))],
+ ([Nil], 'a list),
+ ([Cons], ['a, 'a list] => 'a list),
+ ([list_rec], ['a list, 'b, ['a ,'a list, 'b]=>'b] => 'b),
+ ([list_all], ('a => bool) => ('a list => bool)),
+ ([map], ('a=>'b) => ('a list => 'b list))
+
+\idx{map_def} map(f,xs) == list_rec(xs, Nil, %x l r. Cons(f(x), r)) )
+
+\idx{list_induct}
+ [| P(Nil);
+ !!x xs. [| P(xs) |] ==> P(Cons(x,xs)) |] ==> P(l)
+
+\idx{Cons_not_Nil} ~ Cons(x,xs) = Nil
+\idx{Cons_Cons_eq} (Cons(x,xs)=Cons(y,ys)) = (x=y & xs=ys)
+
+\idx{list_rec_Nil_conv} list_rec(Nil,c,h) = c
+\idx{list_rec_Cons_conv} list_rec(Cons(a,l), c, h) =
+ h(a, l, list_rec(l,c,h))
+
+\idx{map_Nil_conv} map(f,Nil) = Nil
+\idx{map_Cons_conv} map(f, Cons(x,xs)) = Cons(f(x), map(f,xs))
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/LK-eg.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,117 @@
+(**** LK examples -- process using Doc/tout LK-eg.txt ****)
+
+Pretty.setmargin 72; (*existing macros just allow this margin*)
+print_depth 0;
+
+goal LK_Rule.thy "|- EX y. ALL x. P(y)-->P(x)";
+by (resolve_tac [exR] 1);
+by (resolve_tac [allR] 1);
+by (resolve_tac [impR] 1);
+by (resolve_tac [basic] 1);
+(*previous step fails!*)
+by (resolve_tac [exR_thin] 1);
+by (resolve_tac [allR] 1);
+by (resolve_tac [impR] 1);
+by (resolve_tac [basic] 1);
+goal LK_Rule.thy "|- EX y. ALL x. P(y)-->P(x)";
+by (best_tac LK_dup_pack 1);
+
+
+
+goal LK_Rule.thy "|- ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))";
+by (resolve_tac [notR] 1);
+by (resolve_tac [exL] 1);
+by (resolve_tac [allL_thin] 1);
+by (resolve_tac [iffL] 1);
+by (resolve_tac [notL] 2);
+by (resolve_tac [basic] 2);
+by (resolve_tac [notR] 1);
+by (resolve_tac [basic] 1);
+
+
+
+> goal LK_Rule.thy "|- EX y. ALL x. P(y)-->P(x)";
+Level 0
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. |- EX y. ALL x. P(y) --> P(x)
+> by (resolve_tac [exR] 1);
+Level 1
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. |- ALL x. P(?x) --> P(x), EX x. ALL xa. P(x) --> P(xa)
+> by (resolve_tac [allR] 1);
+Level 2
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. !!x. |- P(?x) --> P(x), EX x. ALL xa. P(x) --> P(xa)
+> by (resolve_tac [impR] 1);
+Level 3
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. !!x. P(?x) |- P(x), EX x. ALL xa. P(x) --> P(xa)
+> by (resolve_tac [basic] 1);
+by: tactic failed
+> by (resolve_tac [exR_thin] 1);
+Level 4
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. !!x. P(?x) |- P(x), ALL xa. P(?x7(x)) --> P(xa)
+> by (resolve_tac [allR] 1);
+Level 5
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. !!x xa. P(?x) |- P(x), P(?x7(x)) --> P(xa)
+> by (resolve_tac [impR] 1);
+Level 6
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. !!x xa. P(?x), P(?x7(x)) |- P(x), P(xa)
+> by (resolve_tac [basic] 1);
+Level 7
+ |- EX y. ALL x. P(y) --> P(x)
+No subgoals!
+
+> goal LK_Rule.thy "|- EX y. ALL x. P(y)-->P(x)";
+Level 0
+ |- EX y. ALL x. P(y) --> P(x)
+ 1. |- EX y. ALL x. P(y) --> P(x)
+> by (best_tac LK_dup_pack 1);
+Level 1
+ |- EX y. ALL x. P(y) --> P(x)
+No subgoals!
+
+
+
+
+> goal LK_Rule.thy "|- ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))";
+Level 0
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+> by (resolve_tac [notR] 1);
+Level 1
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. EX x. ALL y. F(y,x) <-> ~F(y,y) |-
+> by (resolve_tac [exL] 1);
+Level 2
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. ALL y. F(y,x) <-> ~F(y,y) |-
+> by (resolve_tac [allL_thin] 1);
+Level 3
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. F(?x2(x),x) <-> ~F(?x2(x),?x2(x)) |-
+> by (resolve_tac [iffL] 1);
+Level 4
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. |- F(?x2(x),x), ~F(?x2(x),?x2(x))
+ 2. !!x. ~F(?x2(x),?x2(x)), F(?x2(x),x) |-
+> by (resolve_tac [notL] 2);
+Level 5
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. |- F(?x2(x),x), ~F(?x2(x),?x2(x))
+ 2. !!x. F(?x2(x),x) |- F(?x2(x),?x2(x))
+> by (resolve_tac [basic] 2);
+Level 6
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. |- F(x,x), ~F(x,x)
+> by (resolve_tac [notR] 1);
+Level 7
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+ 1. !!x. F(x,x) |- F(x,x)
+> by (resolve_tac [basic] 1);
+Level 8
+ |- ~(EX x. ALL y. F(y,x) <-> ~F(y,y))
+No subgoals!
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/LK.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,598 @@
+%% $Id$
+\chapter{First-order sequent calculus}
+The directory~\ttindexbold{LK} implements classical first-order logic through
+Gentzen's sequent calculus (see Gallier~\cite{gallier86} or
+Takeuti~\cite{takeuti87}). Resembling the method of semantic tableaux, the
+calculus is well suited for backwards proof. Assertions have the form
+\(\Gamma\turn \Delta\), where \(\Gamma\) and \(\Delta\) are lists of
+formulae. Associative unification, simulated by higher-order unification,
+handles lists.
+
+The logic is many-sorted, using Isabelle's type classes. The
+class of first-order terms is called {\it term}. No types of individuals
+are provided, but extensions can define types such as $nat::term$ and type
+constructors such as $list::(term)term$. Below, the type variable $\alpha$
+ranges over class {\it term\/}; the equality symbol and quantifiers are
+polymorphic (many-sorted). The type of formulae is~{\it o}, which belongs
+to class {\it logic}.
+
+No generic packages are instantiated, since Isabelle does not provide
+packages for sequent calculi at present. \LK{} implements a classical
+logic theorem prover that is as powerful as the generic classical module,
+except that it does not perform equality reasoning.
+
+
+\section{Unification for lists}
+Higher-order unification includes associative unification as a special
+case, by an encoding that involves function composition
+\cite[page~37]{huet78}. To represent lists, let $C$ be a new constant.
+The empty list is $\lambda x.x$, while $[t@1,t@2,\ldots,t@n]$ is
+represented by
+\[ \lambda x.C(t@1,C(t@2,\ldots,C(t@n,x))). \]
+The unifiers of this with $\lambda x.\Var{f}(\Var{g}(x))$ give all the ways
+of expressing $[t@1,t@2,\ldots,t@n]$ as the concatenation of two lists.
+
+Unlike orthodox associative unification, this technique can represent certain
+infinite sets of unifiers by flex-flex equations. But note that the term
+$\lambda x.C(t,\Var{a})$ does not represent any list. Flex-flex equations
+containing such garbage terms may accumulate during a proof.
+
+This technique lets Isabelle formalize sequent calculus rules,
+where the comma is the associative operator:
+\[ \infer[\conj\hbox{-left}]
+ {\Gamma,P\conj Q,\Delta \turn \Theta}
+ {\Gamma,P,Q,\Delta \turn \Theta} \]
+Multiple unifiers occur whenever this is resolved against a goal containing
+more than one conjunction on the left. Explicit formalization of sequents
+can be tiresome, but gives precise control over contraction and weakening,
+needed to handle relevant and linear logics.
+
+\LK{} exploits this representation of lists. As an alternative, the
+sequent calculus can be formalized using an ordinary representation of
+lists, with a logic program for removing a formula from a list. Amy Felty
+has applied this technique using the language $\lambda$Prolog~\cite{felty91a}.
+
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{Trueprop}& $o\To prop$ & coercion to $prop$ \\
+ \idx{Seqof} & $[o,sobj]\To sobj$ & singleton sequence \\
+ \idx{Not} & $o\To o$ & negation ($\neg$) \\
+ \idx{True} & $o$ & tautology ($\top$) \\
+ \idx{False} & $o$ & absurdity ($\bot$)
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it precedence & \it description \\
+ \idx{ALL} & \idx{All} & $(\alpha\To o)\To o$ & 10 &
+ universal quantifier ($\forall$) \\
+ \idx{EX} & \idx{Ex} & $(\alpha\To o)\To o$ & 10 &
+ existential quantifier ($\exists$) \\
+ \idx{THE} & \idx{The} & $(\alpha\To o)\To \alpha$ & 10 &
+ definite description ($\iota$)
+\end{tabular}
+\end{center}
+\subcaption{Binders}
+
+\begin{center}
+\indexbold{*"=}
+\indexbold{&@{\tt\&}}
+\indexbold{*"|}
+\indexbold{*"-"-">}
+\indexbold{*"<"-">}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt = & $[\alpha,\alpha]\To o$ & Left 50 & equality ($=$) \\
+ \tt \& & $[o,o]\To o$ & Right 35 & conjunction ($\conj$) \\
+ \tt | & $[o,o]\To o$ & Right 30 & disjunction ($\disj$) \\
+ \tt --> & $[o,o]\To o$ & Right 25 & implication ($\imp$) \\
+ \tt <-> & $[o,o]\To o$ & Right 25 & biconditional ($\bimp$)
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+
+\begin{center}
+\begin{tabular}{rrr}
+ \it external & \it internal & \it description \\
+ \tt $\Gamma$ |- $\Delta$ & \tt Trueprop($\Gamma$, $\Delta$) &
+ sequent $\Gamma\turn \Delta$
+\end{tabular}
+\end{center}
+\subcaption{Translations}
+\caption{Syntax of {\tt LK}} \label{lk-syntax}
+\end{figure}
+
+
+\begin{figure}
+\dquotes
+\[\begin{array}{rcl}
+ prop & = & sequence " |- " sequence
+\\[2ex]
+sequence & = & elem \quad (", " elem)^* \\
+ & | & empty
+\\[2ex]
+ elem & = & "\$ " id \\
+ & | & "\$ " var \\
+ & | & formula
+\\[2ex]
+ formula & = & \hbox{expression of type~$o$} \\
+ & | & term " = " term \\
+ & | & "\ttilde\ " formula \\
+ & | & formula " \& " formula \\
+ & | & formula " | " formula \\
+ & | & formula " --> " formula \\
+ & | & formula " <-> " formula \\
+ & | & "ALL~" id~id^* " . " formula \\
+ & | & "EX~~" id~id^* " . " formula \\
+ & | & "THE~" id~ " . " formula
+ \end{array}
+\]
+\caption{Grammar of {\tt LK}} \label{lk-grammar}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{basic} $H, P, $G |- $E, P, $F
+\idx{thinR} $H |- $E, $F ==> $H |- $E, P, $F
+\idx{thinL} $H, $G |- $E ==> $H, P, $G |- $E
+\idx{cut} [| $H |- $E, P; $H, P |- $E |] ==> $H |- $E
+\subcaption{Structural rules}
+
+\idx{refl} $H |- $E, a=a, $F
+\idx{sym} $H |- $E, a=b, $F ==> $H |- $E, b=a, $F
+\idx{trans} [| $H|- $E, a=b, $F; $H|- $E, b=c, $F |] ==>
+ $H|- $E, a=c, $F
+\subcaption{Equality rules}
+
+\idx{True_def} True == False-->False
+\idx{iff_def} P<->Q == (P-->Q) & (Q-->P)
+
+\idx{conjR} [| $H|- $E, P, $F; $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F
+\idx{conjL} $H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E
+
+\idx{disjR} $H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F
+\idx{disjL} [| $H, P, $G |- $E; $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E
+
+\idx{impR} $H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $
+\idx{impL} [| $H,$G |- $E,P; $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E
+
+\idx{notR} $H, P |- $E, $F ==> $H |- $E, ~P, $F
+\idx{notL} $H, $G |- $E, P ==> $H, ~P, $G |- $E
+
+\idx{FalseL} $H, False, $G |- $E
+\subcaption{Propositional rules}
+
+\idx{allR} (!!x.$H|- $E, P(x), $F) ==> $H|- $E, ALL x.P(x), $F
+\idx{allL} $H, P(x), $G, ALL x.P(x) |- $E ==> $H, ALL x.P(x), $G|- $E
+
+\idx{exR} $H|- $E, P(x), $F, EX x.P(x) ==> $H|- $E, EX x.P(x), $F
+\idx{exL} (!!x.$H, P(x), $G|- $E) ==> $H, EX x.P(x), $G|- $E
+
+\idx{The} [| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==>
+ $H |- $E, P(THE x.P(x)), $F
+\subcaption{Quantifier rules}
+\end{ttbox}
+
+\caption{Rules of {\tt LK}} \label{lk-rules}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{conR} $H |- $E, P, $F, P ==> $H |- $E, P, $F
+\idx{conL} $H, P, $G, P |- $E ==> $H, P, $G |- $E
+
+\idx{symL} $H, $G, B = A |- $E ==> $H, A = B, $G |- $E
+
+\idx{TrueR} $H |- $E, True, $F
+
+\idx{iffR} [| $H, P |- $E, Q, $F; $H, Q |- $E, P, $F |] ==>
+ $H |- $E, P<->Q, $F
+
+\idx{iffL} [| $H, $G |- $E, P, Q; $H, Q, P, $G |- $E |] ==>
+ $H, P<->Q, $G |- $E
+
+\idx{allL_thin} $H, P(x), $G |- $E ==> $H, ALL x.P(x), $G |- $E
+\idx{exR_thin} $H |- $E, P(x), $F ==> $H |- $E, EX x.P(x), $F
+\end{ttbox}
+
+\caption{Derived rules for {\tt LK}} \label{lk-derived}
+\end{figure}
+
+
+\section{Syntax and rules of inference}
+Figure~\ref{lk-syntax} gives the syntax for {\tt LK}, which is complicated
+by the representation of sequents. Type $sobj\To sobj$ represents a list
+of formulae.
+
+The {\bf definite description} operator~$\iota x.P(x)$ stands for the
+unique~$a$ satisfying~$P(a)$, if such exists. Since all terms in \LK{}
+denote something, a description is always meaningful, but we do not know
+its value unless $P[x]$ defines it uniquely. The Isabelle notation is
+\hbox{\tt THE $x$.$P[x]$}.
+
+Traditionally, \(\Gamma\) and \(\Delta\) are meta-variables for sequences.
+In Isabelle's notation, the prefix~\verb|$| on a variable makes it range
+over sequences. In a sequent, anything not prefixed by \verb|$| is taken
+as a formula.
+
+The theory has the \ML\ identifier \ttindexbold{LK.thy}.
+Figure~\ref{lk-rules} presents the rules. The connective $\bimp$ is
+defined using $\conj$ and $\imp$. The axiom for basic sequents is
+expressed in a form that provides automatic thinning: redundant formulae
+are simply ignored. The other rules are expressed in the form most
+suitable for backward proof --- they do not require exchange or contraction
+rules. The contraction rules are actually derivable (via cut) in this
+formulation.
+
+Figure~\ref{lk-derived} presents derived rules, including rules for
+$\bimp$. The weakened quantifier rules discard each quantification after a
+single use; in an automatic proof procedure, they guarantee termination,
+but are incomplete. Multiple use of a quantifier can be obtained by a
+contraction rule, which in backward proof duplicates a formula. The tactic
+\ttindex{res_inst_tac} can instantiate the variable~{\tt?P} in these rules,
+specifying the formula to duplicate.
+
+See the files \ttindexbold{LK/lk.thy} and \ttindexbold{LK/lk.ML}
+for complete listings of the rules and derived rules.
+
+
+\section{Tactics for the cut rule}
+According to the cut-elimination theorem, the cut rule can be eliminated
+from proofs of sequents. But the rule is still essential. It can be used
+to structure a proof into lemmas, avoiding repeated proofs of the same
+formula. More importantly, the cut rule can not be eliminated from
+derivations of rules. For example, there is a trivial cut-free proof of
+the sequent \(P\conj Q\turn Q\conj P\).
+Noting this, we might want to derive a rule for swapping the conjuncts
+in a right-hand formula:
+\[ \Gamma\turn \Delta, P\conj Q\over \Gamma\turn \Delta, Q\conj P \]
+The cut rule must be used, for $P\conj Q$ is not a subformula of $Q\conj
+P$. Most cuts directly involve a premise of the rule being derived (a
+meta-assumption). In a few cases, the cut formula is not part of any
+premise, but serves as a bridge between the premises and the conclusion.
+In such proofs, the cut formula is specified by calling an appropriate
+tactic.
+
+\begin{ttbox}
+cutR_tac : string -> int -> tactic
+cutL_tac : string -> int -> tactic
+\end{ttbox}
+These tactics refine a subgoal into two by applying the cut rule. The cut
+formula is given as a string, and replaces some other formula in the sequent.
+\begin{description}
+\item[\ttindexbold{cutR_tac} {\it formula} {\it i}]
+reads an \LK{} formula~$P$, and applies the cut rule to subgoal~$i$. It
+then deletes some formula from the right side of subgoal~$i$, replacing
+that formula by~$P$.
+
+\item[\ttindexbold{cutL_tac} {\it formula} {\it i}]
+reads an \LK{} formula~$P$, and applies the cut rule to subgoal~$i$. It
+then deletes some formula from the let side of the new subgoal $i+1$,
+replacing that formula by~$P$.
+\end{description}
+All the structural rules --- cut, contraction, and thinning --- can be
+applied to particular formulae using \ttindex{res_inst_tac}.
+
+
+\section{Tactics for sequents}
+\begin{ttbox}
+forms_of_seq : term -> term list
+could_res : term * term -> bool
+could_resolve_seq : term * term -> bool
+filseq_resolve_tac : thm list -> int -> int -> tactic
+\end{ttbox}
+Associative unification is not as efficient as it might be, in part because
+the representation of lists defeats some of Isabelle's internal
+optimizations. The following operations implement faster rule application,
+and may have other uses.
+\begin{description}
+\item[\ttindexbold{forms_of_seq} {\it t}]
+returns the list of all formulae in the sequent~$t$, removing sequence
+variables.
+
+\item[\ttindexbold{could_res} $(t,u)$]
+tests whether two formula lists could be resolved. List $t$ is from a
+premise (subgoal), while $u$ is from the conclusion of an object-rule.
+Assuming that each formula in $u$ is surrounded by sequence variables, it
+checks that each conclusion formula is unifiable (using {\tt could_unify})
+with some subgoal formula.
+
+\item[\ttindexbold{could_resolve} $(t,u)$]
+tests whether two sequents could be resolved. Sequent $t$ is a premise
+(subgoal), while $u$ is the conclusion of an object-rule. It uses {\tt
+could_res} to check the left and right sides of the two sequents.
+
+\item[\ttindexbold{filseq_resolve_tac} {\it thms} {\it maxr} {\it i}]
+uses {\tt filter_thms could_resolve} to extract the {\it thms} that are
+applicable to subgoal~$i$. If more than {\it maxr\/} theorems are
+applicable then the tactic fails. Otherwise it calls {\tt resolve_tac}.
+Thus, it is the sequent calculus analogue of \ttindex{filt_resolve_tac}.
+\end{description}
+
+
+
+\section{Packaging sequent rules}
+For theorem proving, rules can be classified as {\bf safe} or {\bf
+unsafe}. An unsafe rule (typically a weakened quantifier rule) may reduce
+a provable goal to an unprovable set of subgoals, and should only be used
+as a last resort.
+
+A {\bf pack} is a pair whose first component is a list of safe
+rules, and whose second is a list of unsafe rules. Packs can be extended
+in an obvious way to allow reasoning with various collections of rules.
+For clarity, \LK{} declares the datatype
+\ttindexbold{pack}:
+\begin{ttbox}
+datatype pack = Pack of thm list * thm list;
+\end{ttbox}
+The contents of any pack can be inspected by pattern-matching. Packs
+support the following operations:
+\begin{ttbox}
+empty_pack : pack
+prop_pack : pack
+LK_pack : pack
+LK_dup_pack : pack
+add_safes : pack * thm list -> pack \hfill{\bf infix 4}
+add_unsafes : pack * thm list -> pack \hfill{\bf infix 4}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{empty_pack}] is the empty pack.
+
+\item[\ttindexbold{prop_pack}] contains the propositional rules, namely
+those for $\conj$, $\disj$, $\neg$, $\imp$ and~$\bimp$, along with the
+rules \ttindex{basic} and \ttindex{refl}. These are all safe.
+
+\item[\ttindexbold{LK_pack}]
+extends {\tt prop_pack} with the safe rules \ttindex{allR}
+and~\ttindex{exL} and the unsafe rules \ttindex{allL_thin} and
+\ttindex{exR_thin}. Search using this is incomplete since quantified
+formulae are used at most once.
+
+\item[\ttindexbold{LK_dup_pack}]
+extends {\tt prop_pack} with the safe rules \ttindex{allR}
+and~\ttindex{exL} and the unsafe rules \ttindex{allL} and~\ttindex{exR}.
+Search using this is complete, since quantified formulae may be reused, but
+frequently fails to terminate. It is generally unsuitable for depth-first
+search.
+
+\item[$pack$ \ttindexbold{add_safes} $rules$]
+adds some safe~$rules$ to the pack~$pack$.
+
+\item[$pack$ \ttindexbold{add_unsafes} $rules$]
+adds some unsafe~$rules$ to the pack~$pack$.
+\end{description}
+
+
+\section{Proof procedures}
+The \LK{} proof procedure is similar to the generic classical module
+described in the {\em Reference Manual}. In fact it is simpler, since it
+works directly with sequents rather than simulating them. There is no need
+to distinguish introduction rules from elimination rules, and of course
+there is no swap rule. As always, Isabelle's classical proof procedures
+are less powerful than resolution theorem provers. But they are more
+natural and flexible, working with an open-ended set of rules.
+
+Backtracking over the choice of a safe rule accomplishes nothing: applying
+them in any order leads to essentially the same result. Backtracking may
+be necessary over basic sequents when they perform unification. Suppose
+that~0, 1, 2,~3 are constants in the subgoals
+\[ \begin{array}{c}
+ P(0), P(1), P(2) \turn P(\Var{a}) \\
+ P(0), P(2), P(3) \turn P(\Var{a}) \\
+ P(1), P(3), P(2) \turn P(\Var{a})
+ \end{array}
+\]
+The only assignment that satisfies all three subgoals is $\Var{a}\mapsto 2$,
+and this can only be discovered by search. The tactics given below permit
+backtracking only over axioms, such as {\tt basic} and {\tt refl}.
+
+
+\subsection{Method A}
+\begin{ttbox}
+reresolve_tac : thm list -> int -> tactic
+repeat_goal_tac : pack -> int -> tactic
+pc_tac : pack -> int -> tactic
+\end{ttbox}
+These tactics use a method developed by Philippe de Groote. A subgoal is
+refined and the resulting subgoals are attempted in reverse order. For
+some reason, this is much faster than attempting the subgoals in order.
+The method is inherently depth-first.
+
+At present, these tactics only work for rules that have no more than two
+premises. They {\bf fail} if they can do nothing.
+\begin{description}
+\item[\ttindexbold{reresolve_tac} $thms$ $i$]
+repeatedly applies the $thms$ to subgoal $i$ and the resulting subgoals.
+
+\item[\ttindexbold{repeat_goal_tac} $pack$ $i$]
+applies the safe rules in the pack to a goal and the resulting subgoals.
+If no safe rule is applicable then it applies an unsafe rule and continues.
+
+\item[\ttindexbold{pc_tac} $pack$ $i$]
+applies {\tt repeat_goal_tac} using depth-first search to solve subgoal~$i$.
+\end{description}
+
+
+\subsection{Method B}
+\begin{ttbox}
+safe_goal_tac : pack -> int -> tactic
+step_tac : pack -> int -> tactic
+fast_tac : pack -> int -> tactic
+best_tac : pack -> int -> tactic
+\end{ttbox}
+These tactics are precisely analogous to those of the generic classical
+module. They use `Method~A' only on safe rules. They {\bf fail} if they
+can do nothing.
+\begin{description}
+\item[\ttindexbold{safe_goal_tac} $pack$ $i$]
+applies the safe rules in the pack to a goal and the resulting subgoals.
+It ignores the unsafe rules.
+
+\item[\ttindexbold{step_tac} $pack$ $i$]
+either applies safe rules (using {\tt safe_goal_tac}) or applies one unsafe
+rule.
+
+\item[\ttindexbold{fast_tac} $pack$ $i$]
+applies {\tt step_tac} using depth-first search to solve subgoal~$i$.
+Despite the names, {\tt pc_tac} is frequently faster.
+
+\item[\ttindexbold{best_tac} $pack$ $i$]
+applies {\tt step_tac} using best-first search to solve subgoal~$i$. It is
+particularly useful for quantifier duplication (using \ttindex{LK_dup_pack}).
+\end{description}
+
+
+
+\section{A simple example of classical reasoning}
+The theorem $\turn\ex{y}\all{x}P(y)\imp P(x)$ is a standard example of the
+classical treatment of the existential quantifier. Classical reasoning
+is easy using~{\LK}, as you can see by comparing this proof with the one
+given in~\S\ref{fol-cla-example}. From a logical point of view, the
+proofs are essentially the same; the key step here is to use \ttindex{exR}
+rather than the weaker~\ttindex{exR_thin}.
+\begin{ttbox}
+goal LK.thy "|- EX y. ALL x. P(y)-->P(x)";
+{\out Level 0}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. |- EX y. ALL x. P(y) --> P(x)}
+by (resolve_tac [exR] 1);
+{\out Level 1}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. |- ALL x. P(?x) --> P(x), EX x. ALL xa. P(x) --> P(xa)}
+\end{ttbox}
+There are now two formulae on the right side. Keeping the existential one
+in reserve, we break down the universal one.
+\begin{ttbox}
+by (resolve_tac [allR] 1);
+{\out Level 2}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. |- P(?x) --> P(x), EX x. ALL xa. P(x) --> P(xa)}
+by (resolve_tac [impR] 1);
+{\out Level 3}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. P(?x) |- P(x), EX x. ALL xa. P(x) --> P(xa)}
+\end{ttbox}
+Because {\LK} is a sequent calculus, the formula~$P(\Var{x})$ does not
+become an assumption; instead, it moves to the left side. The
+resulting subgoal cannot be instantiated to a basic sequent: the bound
+variable~$x$ is not unifiable with the unknown~$\Var{x}$.
+\begin{ttbox}
+by (resolve_tac [basic] 1);
+{\out by: tactic failed}
+\end{ttbox}
+We reuse the existential formula using~\ttindex{exR_thin}, which discards
+it; we will not need it a third time. We again break down the resulting
+formula.
+\begin{ttbox}
+by (resolve_tac [exR_thin] 1);
+{\out Level 4}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x. P(?x) |- P(x), ALL xa. P(?x7(x)) --> P(xa)}
+by (resolve_tac [allR] 1);
+{\out Level 5}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x xa. P(?x) |- P(x), P(?x7(x)) --> P(xa)}
+by (resolve_tac [impR] 1);
+{\out Level 6}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. !!x xa. P(?x), P(?x7(x)) |- P(x), P(xa)}
+\end{ttbox}
+Subgoal~1 seems to offer lots of possibilities. Actually the only useful
+step is instantiating~$\Var{x@7}$ to $\lambda x.x$,
+transforming~$\Var{x@7}(x)$ into~$x$.
+\begin{ttbox}
+by (resolve_tac [basic] 1);
+{\out Level 7}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out No subgoals!}
+\end{ttbox}
+This theorem can be proved automatically. Because it involves quantifier
+duplication, we employ best-first search:
+\begin{ttbox}
+goal LK.thy "|- EX y. ALL x. P(y)-->P(x)";
+{\out Level 0}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out 1. |- EX y. ALL x. P(y) --> P(x)}
+by (best_tac LK_dup_pack 1);
+{\out Level 1}
+{\out |- EX y. ALL x. P(y) --> P(x)}
+{\out No subgoals!}
+\end{ttbox}
+
+
+
+\section{A more complex proof}
+Many of Pelletier's test problems for theorem provers \cite{pelletier86}
+can be solved automatically. Problem~39 concerns set theory, asserting
+that there is no Russell set --- a set consisting of those sets that are
+not members of themselves:
+\[ \turn \neg (\exists x. \forall y. y\in x \bimp y\not\in y) \]
+This does not require special properties of membership; we may
+generalize $x\in y$ to an arbitrary predicate~$F(x,y)$. The theorem has a
+short manual proof. See the directory \ttindexbold{LK/ex} for many more
+examples.
+
+We set the main goal and move the negated formula to the left.
+\begin{ttbox}
+goal LK.thy "|- ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))";
+{\out Level 0}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+by (resolve_tac [notR] 1);
+{\out Level 1}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. EX x. ALL y. F(y,x) <-> ~ F(y,y) |-}
+by (resolve_tac [exL] 1);
+\end{ttbox}
+The right side is empty; we strip both quantifiers from the formula on the
+left.
+\begin{ttbox}
+{\out Level 2}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. ALL y. F(y,x) <-> ~ F(y,y) |-}
+by (resolve_tac [allL_thin] 1);
+{\out Level 3}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. F(?x2(x),x) <-> ~ F(?x2(x),?x2(x)) |-}
+\end{ttbox}
+The rule \ttindex{iffL} says, if $P\bimp Q$ then $P$ and~$Q$ are either
+both true or both false. It yields two subgoals.
+\begin{ttbox}
+by (resolve_tac [iffL] 1);
+{\out Level 4}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. |- F(?x2(x),x), ~ F(?x2(x),?x2(x))}
+{\out 2. !!x. ~ F(?x2(x),?x2(x)), F(?x2(x),x) |-}
+\end{ttbox}
+We must instantiate~$\Var{x@2}$, the shared unknown, to satisfy both
+subgoals. Beginning with subgoal~2, we move a negated formula to the left
+and create a basic sequent.
+\begin{ttbox}
+by (resolve_tac [notL] 2);
+{\out Level 5}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. |- F(?x2(x),x), ~ F(?x2(x),?x2(x))}
+{\out 2. !!x. F(?x2(x),x) |- F(?x2(x),?x2(x))}
+by (resolve_tac [basic] 2);
+{\out Level 6}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. |- F(x,x), ~ F(x,x)}
+\end{ttbox}
+Thanks to the instantiation of~$\Var{x@2}$, subgoal~1 is obviously true.
+\begin{ttbox}
+by (resolve_tac [notR] 1);
+{\out Level 7}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out 1. !!x. F(x,x) |- F(x,x)}
+by (resolve_tac [basic] 1);
+{\out Level 8}
+{\out |- ~ (EX x. ALL y. F(y,x) <-> ~ F(y,y))}
+{\out No subgoals!}
+\end{ttbox}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/Old_HOL.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,1370 @@
+%% $Id$
+\chapter{Higher-order logic}
+The directory~\ttindexbold{HOL} contains a theory for higher-order logic
+based on Gordon's~{\sc hol} system~\cite{gordon88a}, which itself is based on
+Church~\cite{church40}. Andrews~\cite{andrews86} is a full description of
+higher-order logic. Gordon's work has demonstrated that higher-order logic
+is useful for hardware verification; beyond this, it is widely applicable
+in many areas of mathematics. It is weaker than {\ZF} set theory but for
+most applications this does not matter. If you prefer {\ML} to Lisp, you
+will probably prefer {\HOL} to~{\ZF}.
+
+Previous releases of Isabelle included a completely different version
+of~{\HOL}, with explicit type inference rules~\cite{paulson-COLOG}. This
+version no longer exists, but \ttindex{ZF} supports a similar style of
+reasoning.
+
+{\HOL} has a distinct feel, compared with {\ZF} and {\CTT}. It
+identifies object-level types with meta-level types, taking advantage of
+Isabelle's built-in type checker. It identifies object-level functions
+with meta-level functions, so it uses Isabelle's operations for abstraction
+and application. There is no ``apply'' operator: function applications are
+written as simply~$f(a)$ rather than $f{\tt`}a$.
+
+These identifications allow Isabelle to support {\HOL} particularly nicely,
+but they also mean that {\HOL} requires more sophistication from the user
+--- in particular, an understanding of Isabelle's type system. Beginners
+should gain experience by working in first-order logic, before attempting
+to use higher-order logic. This chapter assumes familiarity with~{\FOL{}}.
+
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{Trueprop}& $bool\To prop$ & coercion to $prop$\\
+ \idx{not} & $bool\To bool$ & negation ($\neg$) \\
+ \idx{True} & $bool$ & tautology ($\top$) \\
+ \idx{False} & $bool$ & absurdity ($\bot$) \\
+ \idx{if} & $[bool,\alpha,\alpha]\To\alpha::term$ & conditional \\
+ \idx{Inv} & $(\alpha\To\beta)\To(\beta\To\alpha)$ & function inversion
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\index{"@@{\tt\at}}
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it prec & \it description \\
+ \tt\at & \idx{Eps} & $(\alpha\To bool)\To\alpha::term$ & 10 &
+ Hilbert description ($\epsilon$) \\
+ \idx{!} & \idx{All} & $(\alpha::term\To bool)\To bool$ & 10 &
+ universal quantifier ($\forall$) \\
+ \idx{?} & \idx{Ex} & $(\alpha::term\To bool)\To bool$ & 10 &
+ existential quantifier ($\exists$) \\
+ \idx{?!} & \idx{Ex1} & $(\alpha::term\To bool)\To bool$ & 10 &
+ unique existence ($\exists!$)
+\end{tabular}
+\end{center}
+\subcaption{Binders}
+
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it prec & \it description \\
+ \idx{ALL} & \idx{All} & $(\alpha::term\To bool)\To bool$ & 10 &
+ universal quantifier ($\forall$) \\
+ \idx{EX} & \idx{Ex} & $(\alpha::term\To bool)\To bool$ & 10 &
+ existential quantifier ($\exists$) \\
+ \idx{EX!} & \idx{Ex1} & $(\alpha::term\To bool)\To bool$ & 10 &
+ unique existence ($\exists!$)
+\end{tabular}
+\end{center}
+\subcaption{Alternative quantifiers}
+
+\begin{center}
+\indexbold{*"=}
+\indexbold{&@{\tt\&}}
+\indexbold{*"|}
+\indexbold{*"-"-">}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \idx{o} & $[\beta\To\gamma,\alpha\To\beta]\To (\alpha\To\gamma)$ &
+ Right 50 & composition ($\circ$) \\
+ \tt = & $[\alpha::term,\alpha]\To bool$ & Left 50 & equality ($=$) \\
+ \tt \& & $[bool,bool]\To bool$ & Right 35 & conjunction ($\conj$) \\
+ \tt | & $[bool,bool]\To bool$ & Right 30 & disjunction ($\disj$) \\
+ \tt --> & $[bool,bool]\To bool$ & Right 25 & implication ($\imp$) \\
+ \tt < & $[\alpha::ord,\alpha]\To bool$ & Left 50 & less than ($<$) \\
+ \tt <= & $[\alpha::ord,\alpha]\To bool$ & Left 50 &
+ less than or equals ($\leq$)
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+\caption{Syntax of {\tt HOL}} \label{hol-constants}
+\end{figure}
+
+
+\begin{figure}
+\dquotes
+\[\begin{array}{rcl}
+ term & = & \hbox{expression of class~$term$} \\
+ & | & "\at~~" id~id^* " . " formula \\[2ex]
+ formula & = & \hbox{expression of type~$bool$} \\
+ & | & term " = " term \\
+ & | & term " \ttilde= " term \\
+ & | & term " < " term \\
+ & | & term " <= " term \\
+ & | & "\ttilde\ " formula \\
+ & | & formula " \& " formula \\
+ & | & formula " | " formula \\
+ & | & formula " --> " formula \\
+ & | & "!~~~" id~id^* " . " formula \\
+ & | & "?~~~" id~id^* " . " formula \\
+ & | & "?!~~" id~id^* " . " formula \\
+ & | & "ALL~" id~id^* " . " formula \\
+ & | & "EX~~" id~id^* " . " formula \\
+ & | & "EX!~" id~id^* " . " formula
+ \end{array}
+\]
+\subcaption{Grammar}
+\caption{Full grammar for {\HOL}} \label{hol-grammar}
+\end{figure}
+
+
+\section{Syntax}
+Type inference is automatic, exploiting Isabelle's type classes. The class
+of higher-order terms is called {\it term\/}; type variables range over
+this class by default. The equality symbol and quantifiers are polymorphic
+over class {\it term}.
+
+Class {\it ord\/} consists of all ordered types; the relations $<$ and
+$\leq$ are polymorphic over this class, as are the functions
+\ttindex{mono}, \ttindex{min} and \ttindex{max}. Three other
+type classes --- {\it plus}, {\it minus\/} and {\it times} --- permit
+overloading of the operators {\tt+}, {\tt-} and {\tt*}. In particular,
+{\tt-} is overloaded for set difference and subtraction.
+\index{*"+}\index{-@{\tt-}}\index{*@{\tt*}}
+
+Figure~\ref{hol-constants} lists the constants (including infixes and
+binders), while Figure~\ref{hol-grammar} presents the grammar. Note that
+$a$\verb|~=|$b$ is translated to \verb|~(|$a$=$b$\verb|)|.
+
+\subsection{Types}
+The type of formulae, {\it bool} belongs to class {\it term}; thus,
+formulae are terms. The built-in type~$fun$, which constructs function
+types, is overloaded such that $\sigma\To\tau$ belongs to class~$term$ if
+$\sigma$ and~$\tau$ do; this allows quantification over functions. Types
+in {\HOL} must be non-empty because of the form of quantifier
+rules~\cite[\S7]{paulson-COLOG}.
+
+Gordon's {\sc hol} system supports {\bf type definitions}. A type is
+defined by exhibiting an existing type~$\sigma$, a predicate~$P::\sigma\To
+bool$, and a theorem of the form $\exists x::\sigma.P(x)$. Thus~$P$
+specifies a non-empty subset of~$\sigma$, and the new type denotes this
+subset. New function constants are generated to establish an isomorphism
+between the new type and the subset. If type~$\sigma$ involves type
+variables $\alpha@1$, \ldots, $\alpha@n$, then the type definition creates
+a type constructor $(\alpha@1,\ldots,\alpha@n)ty$ rather than a particular
+type.
+
+Isabelle does not support type definitions at present. Instead, they are
+mimicked by explicit definitions of isomorphism functions. These should be
+accompanied by theorems of the form $\exists x::\sigma.P(x)$, but this is
+not checked.
+
+
+\subsection{Binders}
+Hilbert's {\bf description} operator~$\epsilon x.P[x]$ stands for some~$a$
+satisfying~$P[a]$, if such exists. Since all terms in {\HOL} denote
+something, a description is always meaningful, but we do not know its value
+unless $P[x]$ defines it uniquely. We may write descriptions as
+\ttindexbold{Eps}($P$) or use the syntax
+\hbox{\tt \at $x$.$P[x]$}. Existential quantification is defined
+by
+\[ \exists x.P(x) \equiv P(\epsilon x.P(x)) \]
+The unique existence quantifier, $\exists!x.P[x]$, is defined in terms
+of~$\exists$ and~$\forall$. An Isabelle binder, it admits nested
+quantifications. For instance, $\exists!x y.P(x,y)$ abbreviates
+$\exists!x. \exists!y.P(x,y)$; note that this does not mean that there
+exists a unique pair $(x,y)$ satisfying~$P(x,y)$.
+
+\index{*"!}\index{*"?}
+Quantifiers have two notations. As in Gordon's {\sc hol} system, {\HOL}
+uses~{\tt!}\ and~{\tt?}\ to stand for $\forall$ and $\exists$. The
+existential quantifier must be followed by a space; thus {\tt?x} is an
+unknown, while \verb'? x.f(x)=y' is a quantification. Isabelle's usual
+notation for quantifiers, \ttindex{ALL} and \ttindex{EX}, is also
+available. Both notations are accepted for input. The {\ML} reference
+\ttindexbold{HOL_quantifiers} governs the output notation. If set to {\tt
+true}, then~{\tt!}\ and~{\tt?}\ are displayed; this is the default. If set
+to {\tt false}, then~{\tt ALL} and~{\tt EX} are displayed.
+
+\begin{warn}
+Although the description operator does not usually allow iteration of the
+form \hbox{\tt \at $x@1 \dots x@n$.$P[x@1,\dots,x@n]$}, there are cases where
+this is legal. If \hbox{\tt \at $y$.$P[x,y]$} is of type~{\it bool},
+then \hbox{\tt \at $x\,y$.$P[x,y]$} is legal. The pretty printer will
+display \hbox{\tt \at $x$.\at $y$.$P[x,y]$} as
+\hbox{\tt \at $x\,y$.$P[x,y]$}.
+\end{warn}
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{refl} t = t::'a
+\idx{subst} [| s=t; P(s) |] ==> P(t::'a)
+\idx{ext} (!!x::'a. f(x)::'b = g(x)) ==> (%x.f(x)) = (%x.g(x))
+\idx{impI} (P ==> Q) ==> P-->Q
+\idx{mp} [| P-->Q; P |] ==> Q
+\idx{iff} (P-->Q) --> (Q-->P) --> (P=Q)
+\idx{selectI} P(x::'a) ==> P(@x.P(x))
+\idx{True_or_False} (P=True) | (P=False)
+\subcaption{basic rules}
+
+\idx{True_def} True = ((%x.x)=(%x.x))
+\idx{All_def} All = (%P. P = (%x.True))
+\idx{Ex_def} Ex = (%P. P(@x.P(x)))
+\idx{False_def} False = (!P.P)
+\idx{not_def} not = (%P. P-->False)
+\idx{and_def} op & = (%P Q. !R. (P-->Q-->R) --> R)
+\idx{or_def} op | = (%P Q. !R. (P-->R) --> (Q-->R) --> R)
+\idx{Ex1_def} Ex1 = (%P. ? x. P(x) & (! y. P(y) --> y=x))
+\subcaption{Definitions of the logical constants}
+
+\idx{Inv_def} Inv = (%(f::'a=>'b) y. @x. f(x)=y)
+\idx{o_def} op o = (%(f::'b=>'c) g (x::'a). f(g(x)))
+\idx{if_def} if = (%P x y.@z::'a.(P=True --> z=x) & (P=False --> z=y))
+\subcaption{Further definitions}
+\end{ttbox}
+\caption{Rules of {\tt HOL}} \label{hol-rules}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{sym} s=t ==> t=s
+\idx{trans} [| r=s; s=t |] ==> r=t
+\idx{ssubst} [| t=s; P(s) |] ==> P(t::'a)
+\idx{box_equals} [| a=b; a=c; b=d |] ==> c=d
+\idx{arg_cong} s=t ==> f(s)=f(t)
+\idx{fun_cong} s::'a=>'b = t ==> s(x)=t(x)
+\subcaption{Equality}
+
+\idx{TrueI} True
+\idx{FalseE} False ==> P
+
+\idx{conjI} [| P; Q |] ==> P&Q
+\idx{conjunct1} [| P&Q |] ==> P
+\idx{conjunct2} [| P&Q |] ==> Q
+\idx{conjE} [| P&Q; [| P; Q |] ==> R |] ==> R
+
+\idx{disjI1} P ==> P|Q
+\idx{disjI2} Q ==> P|Q
+\idx{disjE} [| P | Q; P ==> R; Q ==> R |] ==> R
+
+\idx{notI} (P ==> False) ==> ~ P
+\idx{notE} [| ~ P; P |] ==> R
+\idx{impE} [| P-->Q; P; Q ==> R |] ==> R
+\subcaption{Propositional logic}
+
+\idx{iffI} [| P ==> Q; Q ==> P |] ==> P=Q
+\idx{iffD1} [| P=Q; P |] ==> Q
+\idx{iffD2} [| P=Q; Q |] ==> P
+\idx{iffE} [| P=Q; [| P --> Q; Q --> P |] ==> R |] ==> R
+
+\idx{eqTrueI} P ==> P=True
+\idx{eqTrueE} P=True ==> P
+\subcaption{Logical equivalence}
+
+\end{ttbox}
+\caption{Derived rules for {\HOL}} \label{hol-lemmas1}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{allI} (!!x::'a. P(x)) ==> !x. P(x)
+\idx{spec} !x::'a.P(x) ==> P(x)
+\idx{allE} [| !x.P(x); P(x) ==> R |] ==> R
+\idx{all_dupE} [| !x.P(x); [| P(x); !x.P(x) |] ==> R |] ==> R
+
+\idx{exI} P(x) ==> ? x::'a.P(x)
+\idx{exE} [| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q
+
+\idx{ex1I} [| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)
+\idx{ex1E} [| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R
+ |] ==> R
+
+\idx{select_equality} [| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a
+\subcaption{Quantifiers and descriptions}
+
+\idx{ccontr} (~P ==> False) ==> P
+\idx{classical} (~P ==> P) ==> P
+\idx{excluded_middle} ~P | P
+
+\idx{disjCI} (~Q ==> P) ==> P|Q
+\idx{exCI} (! x. ~ P(x) ==> P(a)) ==> ? x.P(x)
+\idx{impCE} [| P-->Q; ~ P ==> R; Q ==> R |] ==> R
+\idx{iffCE} [| P<->Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R
+\idx{notnotD} ~~P ==> P
+\idx{swap} ~P ==> (~Q ==> P) ==> Q
+\subcaption{Classical logic}
+
+\idx{if_True} if(True,x,y) = x
+\idx{if_False} if(False,x,y) = y
+\idx{if_P} P ==> if(P,x,y) = x
+\idx{if_not_P} ~ P ==> if(P,x,y) = y
+\idx{expand_if} P(if(Q,x,y)) = ((Q --> P(x)) & (~Q --> P(y)))
+\subcaption{Conditionals}
+\end{ttbox}
+\caption{More derived rules} \label{hol-lemmas2}
+\end{figure}
+
+
+\section{Rules of inference}
+The basic theory has the {\ML} identifier \ttindexbold{HOL.thy}. However,
+many further theories are defined, introducing product and sum types, the
+natural numbers, etc.
+
+Figure~\ref{hol-rules} shows the inference rules with their~{\ML} names.
+They follow standard practice in higher-order logic: only a few connectives
+are taken as primitive, with the remainder defined obscurely.
+
+Unusually, the definitions use object-equality~({\tt=}) rather than
+meta-equality~({\tt==}). This is possible because equality in higher-order
+logic may equate formulae and even functions over formulae. On the other
+hand, meta-equality is Isabelle's usual symbol for making definitions.
+Take care to note which form of equality is used before attempting a proof.
+
+Some of the rules mention type variables; for example, {\tt refl} mentions
+the type variable~{\tt'a}. This facilitates explicit instantiation of the
+type variables. By default, such variables range over class {\it term}.
+
+\begin{warn}
+Where function types are involved, Isabelle's unification code does not
+guarantee to find instantiations for type variables automatically. If
+resolution fails for no obvious reason, try setting \ttindex{show_types} to
+{\tt true}, causing Isabelle to display types of terms. (Possibly, set
+\ttindex{show_sorts} to {\tt true} also, causing Isabelle to display sorts.)
+Be prepared to use \ttindex{res_inst_tac} instead of {\tt resolve_tac}.
+Setting \ttindex{Unify.trace_types} to {\tt true} causes Isabelle to
+report omitted search paths during unification.
+\end{warn}
+
+Here are further comments on the rules:
+\begin{description}
+\item[\ttindexbold{ext}]
+expresses extensionality of functions.
+\item[\ttindexbold{iff}]
+asserts that logically equivalent formulae are equal.
+\item[\ttindexbold{selectI}]
+gives the defining property of the Hilbert $\epsilon$-operator. The
+derived rule \ttindexbold{select_equality} (see below) is often easier to use.
+\item[\ttindexbold{True_or_False}]
+makes the logic classical.\footnote{In fact, the $\epsilon$-operator
+already makes the logic classical, as shown by Diaconescu;
+see Paulson~\cite{paulson-COLOG} for details.}
+\end{description}
+
+\begin{warn}
+{\HOL} has no if-and-only-if connective; logical equivalence is expressed
+using equality. But equality has a high precedence, as befitting a
+relation, while if-and-only-if typically has the lowest precedence. Thus,
+$\neg\neg P=P$ abbreviates $\neg\neg (P=P)$ and not $(\neg\neg P)=P$. When
+using $=$ to mean logical equivalence, enclose both operands in
+parentheses.
+\end{warn}
+
+Some derived rules are shown in Figures~\ref{hol-lemmas1}
+and~\ref{hol-lemmas2}, with their {\ML} names. These include natural rules
+for the logical connectives, as well as sequent-style elimination rules for
+conjunctions, implications, and universal quantifiers.
+
+Note the equality rules: \ttindexbold{ssubst} performs substitution in
+backward proofs, while \ttindexbold{box_equals} supports reasoning by
+simplifying both sides of an equation.
+
+See the files \ttindexbold{HOL/hol.thy} and
+\ttindexbold{HOL/hol.ML} for complete listings of the rules and
+derived rules.
+
+
+\section{Generic packages}
+{\HOL} instantiates most of Isabelle's generic packages;
+see \ttindexbold{HOL/ROOT.ML} for details.
+\begin{itemize}
+\item
+Because it includes a general substitution rule, {\HOL} instantiates the
+tactic \ttindex{hyp_subst_tac}, which substitutes for an equality
+throughout a subgoal and its hypotheses.
+\item
+It instantiates the simplifier, defining~\ttindexbold{HOL_ss} as the
+simplification set for higher-order logic. Equality~($=$), which also
+expresses logical equivalence, may be used for rewriting. See the file
+\ttindexbold{HOL/simpdata.ML} for a complete listing of the simplification
+rules.
+\item
+It instantiates the classical reasoning module. See~\S\ref{hol-cla-prover}
+for details.
+\end{itemize}
+
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+\index{"{"}@{\tt\{\}}}
+ {\tt\{\}} & $\alpha\,set$ & the empty set \\
+ \idx{insert} & $[\alpha,\alpha\,set]\To \alpha\,set$
+ & insertion of element \\
+ \idx{Collect} & $(\alpha\To bool)\To\alpha\,set$
+ & comprehension \\
+ \idx{Compl} & $(\alpha\,set)\To\alpha\,set$
+ & complement \\
+ \idx{INTER} & $[\alpha\,set,\alpha\To\beta\,set]\To\beta\,set$
+ & intersection over a set\\
+ \idx{UNION} & $[\alpha\,set,\alpha\To\beta\,set]\To\beta\,set$
+ & union over a set\\
+ \idx{Inter} & $((\alpha\,set)set)\To\alpha\,set$
+ &set of sets intersection \\
+ \idx{Union} & $((\alpha\,set)set)\To\alpha\,set$
+ &set of sets union \\
+ \idx{range} & $(\alpha\To\beta )\To\beta\,set$
+ & range of a function \\[1ex]
+ \idx{Ball}~~\idx{Bex} & $[\alpha\,set,\alpha\To bool]\To bool$
+ & bounded quantifiers \\
+ \idx{mono} & $(\alpha\,set\To\beta\,set)\To bool$
+ & monotonicity \\
+ \idx{inj}~~\idx{surj}& $(\alpha\To\beta )\To bool$
+ & injective/surjective \\
+ \idx{inj_onto} & $[\alpha\To\beta ,\alpha\,set]\To bool$
+ & injective over subset
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{center}
+\begin{tabular}{llrrr}
+ \it symbol &\it name &\it meta-type & \it prec & \it description \\
+ \idx{INT} & \idx{INTER1} & $(\alpha\To\beta\,set)\To\beta\,set$ & 10 &
+ intersection over a type\\
+ \idx{UN} & \idx{UNION1} & $(\alpha\To\beta\,set)\To\beta\,set$ & 10 &
+ union over a type
+\end{tabular}
+\end{center}
+\subcaption{Binders}
+
+\begin{center}
+\indexbold{*"`"`}
+\indexbold{*":}
+\indexbold{*"<"=}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt `` & $[\alpha\To\beta ,\alpha\,set]\To (\beta\,set)$
+ & Left 90 & image \\
+ \idx{Int} & $[\alpha\,set,\alpha\,set]\To\alpha\,set$
+ & Left 70 & intersection ($\inter$) \\
+ \idx{Un} & $[\alpha\,set,\alpha\,set]\To\alpha\,set$
+ & Left 65 & union ($\union$) \\
+ \tt: & $[\alpha ,\alpha\,set]\To bool$
+ & Left 50 & membership ($\in$) \\
+ \tt <= & $[\alpha\,set,\alpha\,set]\To bool$
+ & Left 50 & subset ($\subseteq$)
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+\caption{Syntax of {\HOL}'s set theory} \label{hol-set-syntax}
+\end{figure}
+
+
+\begin{figure}
+\begin{center} \tt\frenchspacing
+\begin{tabular}{rrr}
+ \it external & \it internal & \it description \\
+ $a$ \ttilde: $b$ & \ttilde($a$ : $b$) & \rm negated membership\\
+ \{$a@1$, $\ldots$, $a@n$\} & insert($a@1$,$\cdots$,insert($a@n$,0)) &
+ \rm finite set \\
+ \{$x$.$P[x]$\} & Collect($\lambda x.P[x]$) &
+ \rm comprehension \\
+ \idx{INT} $x$:$A$.$B[x]$ & INTER($A$,$\lambda x.B[x]$) &
+ \rm intersection over a set \\
+ \idx{UN} $x$:$A$.$B[x]$ & UNION($A$,$\lambda x.B[x]$) &
+ \rm union over a set \\
+ \idx{!} $x$:$A$.$P[x]$ & Ball($A$,$\lambda x.P[x]$) &
+ \rm bounded $\forall$ \\
+ \idx{?} $x$:$A$.$P[x]$ & Bex($A$,$\lambda x.P[x]$) &
+ \rm bounded $\exists$ \\[1ex]
+ \idx{ALL} $x$:$A$.$P[x]$ & Ball($A$,$\lambda x.P[x]$) &
+ \rm bounded $\forall$ \\
+ \idx{EX} $x$:$A$.$P[x]$ & Bex($A$,$\lambda x.P[x]$) &
+ \rm bounded $\exists$
+\end{tabular}
+\end{center}
+\subcaption{Translations}
+
+\dquotes
+\[\begin{array}{rcl}
+ term & = & \hbox{other terms\ldots} \\
+ & | & "\{\}" \\
+ & | & "\{ " term\; ("," term)^* " \}" \\
+ & | & "\{ " id " . " formula " \}" \\
+ & | & term " `` " term \\
+ & | & term " Int " term \\
+ & | & term " Un " term \\
+ & | & "INT~~" id ":" term " . " term \\
+ & | & "UN~~~" id ":" term " . " term \\
+ & | & "INT~~" id~id^* " . " term \\
+ & | & "UN~~~" id~id^* " . " term \\[2ex]
+ formula & = & \hbox{other formulae\ldots} \\
+ & | & term " : " term \\
+ & | & term " \ttilde: " term \\
+ & | & term " <= " term \\
+ & | & "!~~~" id ":" term " . " formula \\
+ & | & "?~~~" id ":" term " . " formula \\
+ & | & "ALL " id ":" term " . " formula \\
+ & | & "EX~~" id ":" term " . " formula
+ \end{array}
+\]
+\subcaption{Full Grammar}
+\caption{Syntax of {\HOL}'s set theory (continued)} \label{hol-set-syntax2}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{mem_Collect_eq} (a : \{x.P(x)\}) = P(a)
+\idx{Collect_mem_eq} \{x.x:A\} = A
+\subcaption{Isomorphisms between predicates and sets}
+
+\idx{empty_def} \{\} == \{x.x= False\}
+\idx{insert_def} insert(a,B) == \{x.x=a\} Un B
+\idx{Ball_def} Ball(A,P) == ! x. x:A --> P(x)
+\idx{Bex_def} Bex(A,P) == ? x. x:A & P(x)
+\idx{subset_def} A <= B == ! x:A. x:B
+\idx{Un_def} A Un B == \{x.x:A | x:B\}
+\idx{Int_def} A Int B == \{x.x:A & x:B\}
+\idx{set_diff_def} A - B == \{x.x:A & x~:B\}
+\idx{Compl_def} Compl(A) == \{x. ~ x:A\}
+\idx{INTER_def} INTER(A,B) == \{y. ! x:A. y: B(x)\}
+\idx{UNION_def} UNION(A,B) == \{y. ? x:A. y: B(x)\}
+\idx{INTER1_def} INTER1(B) == INTER(\{x.True\}, B)
+\idx{UNION1_def} UNION1(B) == UNION(\{x.True\}, B)
+\idx{Inter_def} Inter(S) == (INT x:S. x)
+\idx{Union_def} Union(S) == (UN x:S. x)
+\idx{image_def} f``A == \{y. ? x:A. y=f(x)\}
+\idx{range_def} range(f) == \{y. ? x. y=f(x)\}
+\idx{mono_def} mono(f) == !A B. A <= B --> f(A) <= f(B)
+\idx{inj_def} inj(f) == ! x y. f(x)=f(y) --> x=y
+\idx{surj_def} surj(f) == ! y. ? x. y=f(x)
+\idx{inj_onto_def} inj_onto(f,A) == !x:A. !y:A. f(x)=f(y) --> x=y
+\subcaption{Definitions}
+\end{ttbox}
+\caption{Rules of {\HOL}'s set theory} \label{hol-set-rules}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{CollectI} [| P(a) |] ==> a : \{x.P(x)\}
+\idx{CollectD} [| a : \{x.P(x)\} |] ==> P(a)
+\idx{CollectE} [| a : \{x.P(x)\}; P(a) ==> W |] ==> W
+\idx{Collect_cong} [| !!x. P(x)=Q(x) |] ==> \{x. P(x)\} = \{x. Q(x)\}
+\subcaption{Comprehension}
+
+\idx{ballI} [| !!x. x:A ==> P(x) |] ==> ! x:A. P(x)
+\idx{bspec} [| ! x:A. P(x); x:A |] ==> P(x)
+\idx{ballE} [| ! x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q
+\idx{ball_cong} [| A=A'; !!x. x:A' ==> P(x) = P'(x) |] ==>
+ (! x:A. P(x)) = (! x:A'. P'(x))
+
+\idx{bexI} [| P(x); x:A |] ==> ? x:A. P(x)
+\idx{bexCI} [| ! x:A. ~ P(x) ==> P(a); a:A |] ==> ? x:A.P(x)
+\idx{bexE} [| ? x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q
+\subcaption{Bounded quantifiers}
+
+\idx{subsetI} (!!x.x:A ==> x:B) ==> A <= B
+\idx{subsetD} [| A <= B; c:A |] ==> c:B
+\idx{subsetCE} [| A <= B; ~ (c:A) ==> P; c:B ==> P |] ==> P
+
+\idx{subset_refl} A <= A
+\idx{subset_antisym} [| A <= B; B <= A |] ==> A = B
+\idx{subset_trans} [| A<=B; B<=C |] ==> A<=C
+
+\idx{set_ext} [| !!x. (x:A) = (x:B) |] ==> A = B
+\idx{equalityD1} A = B ==> A<=B
+\idx{equalityD2} A = B ==> B<=A
+\idx{equalityE} [| A = B; [| A<=B; B<=A |] ==> P |] ==> P
+
+\idx{equalityCE} [| A = B; [| c:A; c:B |] ==> P;
+ [| ~ c:A; ~ c:B |] ==> P
+ |] ==> P
+\subcaption{The subset and equality relations}
+\end{ttbox}
+\caption{Derived rules for set theory} \label{hol-set1}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{emptyE} a : \{\} ==> P
+
+\idx{insertI1} a : insert(a,B)
+\idx{insertI2} a : B ==> a : insert(b,B)
+\idx{insertE} [| a : insert(b,A); a=b ==> P; a:A ==> P |] ==>
+
+\idx{ComplI} [| c:A ==> False |] ==> c : Compl(A)
+\idx{ComplD} [| c : Compl(A) |] ==> ~ c:A
+
+\idx{UnI1} c:A ==> c : A Un B
+\idx{UnI2} c:B ==> c : A Un B
+\idx{UnCI} (~c:B ==> c:A) ==> c : A Un B
+\idx{UnE} [| c : A Un B; c:A ==> P; c:B ==> P |] ==> P
+
+\idx{IntI} [| c:A; c:B |] ==> c : A Int B
+\idx{IntD1} c : A Int B ==> c:A
+\idx{IntD2} c : A Int B ==> c:B
+\idx{IntE} [| c : A Int B; [| c:A; c:B |] ==> P |] ==> P
+
+\idx{UN_I} [| a:A; b: B(a) |] ==> b: (UN x:A. B(x))
+\idx{UN_E} [| b: (UN x:A. B(x)); !!x.[| x:A; b:B(x) |] ==> R |] ==> R
+
+\idx{INT_I} (!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))
+\idx{INT_D} [| b: (INT x:A. B(x)); a:A |] ==> b: B(a)
+\idx{INT_E} [| b: (INT x:A. B(x)); b: B(a) ==> R; ~ a:A ==> R |] ==> R
+
+\idx{UnionI} [| X:C; A:X |] ==> A : Union(C)
+\idx{UnionE} [| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R
+
+\idx{InterI} [| !!X. X:C ==> A:X |] ==> A : Inter(C)
+\idx{InterD} [| A : Inter(C); X:C |] ==> A:X
+\idx{InterE} [| A : Inter(C); A:X ==> R; ~ X:C ==> R |] ==> R
+\end{ttbox}
+\caption{Further derived rules for set theory} \label{hol-set2}
+\end{figure}
+
+
+\section{A formulation of set theory}
+Historically, higher-order logic gives a foundation for Russell and
+Whitehead's theory of classes. Let us use modern terminology and call them
+{\bf sets}, but note that these sets are distinct from those of {\ZF} set
+theory, and behave more like {\ZF} classes.
+\begin{itemize}
+\item
+Sets are given by predicates over some type~$\sigma$. Types serve to
+define universes for sets, but type checking is still significant.
+\item
+There is a universal set (for each type). Thus, sets have complements, and
+may be defined by absolute comprehension.
+\item
+Although sets may contain other sets as elements, the containing set must
+have a more complex type.
+\end{itemize}
+Finite unions and intersections have the same behaviour in {\HOL} as they
+do in~{\ZF}. In {\HOL} the intersection of the empty set is well-defined,
+denoting the universal set for the given type.
+
+\subsection{Syntax of set theory}
+The type $\alpha\,set$ is essentially the same as $\alpha\To bool$. The new
+type is defined for clarity and to avoid complications involving function
+types in unification. Since Isabelle does not support type definitions (as
+discussed above), the isomorphisms between the two types are declared
+explicitly. Here they are natural: {\tt Collect} maps $\alpha\To bool$ to
+$\alpha\,set$, while \hbox{\tt op :} maps in the other direction (ignoring
+argument order).
+
+Figure~\ref{hol-set-syntax} lists the constants, infixes, and syntax
+translations. Figure~\ref{hol-set-syntax2} presents the grammar of the new
+constructs. Infix operators include union and intersection ($A\union B$
+and $A\inter B$), the subset and membership relations, and the image
+operator~{\tt``}. Note that $a$\verb|~:|$b$ is translated to
+\verb|~(|$a$:$b$\verb|)|. The {\tt\{\ldots\}} notation abbreviates finite
+sets constructed in the obvious manner using~{\tt insert} and~$\{\}$ (the
+empty set):
+\begin{eqnarray*}
+ \{a,b,c\} & \equiv & {\tt insert}(a,{\tt insert}(b,{\tt insert}(c,\emptyset)))
+\end{eqnarray*}
+
+The set \hbox{\tt\{$x$.$P[x]$\}} consists of all $x$ (of suitable type)
+that satisfy~$P[x]$, where $P[x]$ is a formula that may contain free
+occurrences of~$x$. This syntax expands to \ttindexbold{Collect}$(\lambda
+x.P[x])$.
+
+The set theory defines two {\bf bounded quantifiers}:
+\begin{eqnarray*}
+ \forall x\in A.P[x] &\hbox{which abbreviates}& \forall x. x\in A\imp P[x] \\
+ \exists x\in A.P[x] &\hbox{which abbreviates}& \exists x. x\in A\conj P[x]
+\end{eqnarray*}
+The constants~\ttindexbold{Ball} and~\ttindexbold{Bex} are defined
+accordingly. Instead of {\tt Ball($A$,$P$)} and {\tt Bex($A$,$P$)} we may
+write\index{*"!}\index{*"?}\index{*ALL}\index{*EX}
+\hbox{\tt !~$x$:$A$.$P[x]$} and \hbox{\tt ?~$x$:$A$.$P[x]$}.
+Isabelle's usual notation, \ttindex{ALL} and \ttindex{EX}, is also
+available. As with
+ordinary quantifiers, the contents of \ttindexbold{HOL_quantifiers} specifies
+which notation should be used for output.
+
+Unions and intersections over sets, namely $\bigcup@{x\in A}B[x]$ and
+$\bigcap@{x\in A}B[x]$, are written
+\ttindexbold{UN}~\hbox{\tt$x$:$A$.$B[x]$} and
+\ttindexbold{INT}~\hbox{\tt$x$:$A$.$B[x]$}.
+Unions and intersections over types, namely $\bigcup@x B[x]$ and
+$\bigcap@x B[x]$, are written
+\ttindexbold{UN}~\hbox{\tt$x$.$B[x]$} and
+\ttindexbold{INT}~\hbox{\tt$x$.$B[x]$}; they are equivalent to the previous
+union/intersection operators when $A$ is the universal set.
+The set of set union and intersection operators ($\bigcup A$ and $\bigcap
+A$) are not binders, but equals $\bigcup@{x\in A}x$ and $\bigcap@{x\in
+ A}x$, respectively.
+
+\subsection{Axioms and rules of set theory}
+The axioms \ttindexbold{mem_Collect_eq} and
+\ttindexbold{Collect_mem_eq} assert that the functions {\tt Collect} and
+\hbox{\tt op :} are isomorphisms.
+All the other axioms are definitions; see Figure~\ref{hol-set-rules}.
+These include straightforward properties of functions: image~({\tt``}) and
+{\tt range}, and predicates concerning monotonicity, injectiveness, etc.
+
+{\HOL}'s set theory has the {\ML} identifier \ttindexbold{Set.thy}.
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{imageI} [| x:A |] ==> f(x) : f``A
+\idx{imageE} [| b : f``A; !!x.[| b=f(x); x:A |] ==> P |] ==> P
+
+\idx{rangeI} f(x) : range(f)
+\idx{rangeE} [| b : range(f); !!x.[| b=f(x) |] ==> P |] ==> P
+
+\idx{monoI} [| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)
+\idx{monoD} [| mono(f); A <= B |] ==> f(A) <= f(B)
+
+\idx{injI} [| !! x y. f(x) = f(y) ==> x=y |] ==> inj(f)
+\idx{inj_inverseI} (!!x. g(f(x)) = x) ==> inj(f)
+\idx{injD} [| inj(f); f(x) = f(y) |] ==> x=y
+
+\idx{Inv_f_f} inj(f) ==> Inv(f,f(x)) = x
+\idx{f_Inv_f} y : range(f) ==> f(Inv(f,y)) = y
+
+\idx{Inv_injective}
+ [| Inv(f,x)=Inv(f,y); x: range(f); y: range(f) |] ==> x=y
+
+\idx{inj_ontoI}
+ (!! x y. [| f(x) = f(y); x:A; y:A |] ==> x=y) ==> inj_onto(f,A)
+
+\idx{inj_onto_inverseI}
+ (!!x. x:A ==> g(f(x)) = x) ==> inj_onto(f,A)
+
+\idx{inj_ontoD}
+ [| inj_onto(f,A); f(x)=f(y); x:A; y:A |] ==> x=y
+
+\idx{inj_onto_contraD}
+ [| inj_onto(f,A); x~=y; x:A; y:A |] ==> ~ f(x)=f(y)
+\end{ttbox}
+\caption{Derived rules involving functions} \label{hol-fun}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{Union_upper} B:A ==> B <= Union(A)
+\idx{Union_least} [| !!X. X:A ==> X<=C |] ==> Union(A) <= C
+
+\idx{Inter_lower} B:A ==> Inter(A) <= B
+\idx{Inter_greatest} [| !!X. X:A ==> C<=X |] ==> C <= Inter(A)
+
+\idx{Un_upper1} A <= A Un B
+\idx{Un_upper2} B <= A Un B
+\idx{Un_least} [| A<=C; B<=C |] ==> A Un B <= C
+
+\idx{Int_lower1} A Int B <= A
+\idx{Int_lower2} A Int B <= B
+\idx{Int_greatest} [| C<=A; C<=B |] ==> C <= A Int B
+\end{ttbox}
+\caption{Derived rules involving subsets} \label{hol-subset}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{Int_absorb} A Int A = A
+\idx{Int_commute} A Int B = B Int A
+\idx{Int_assoc} (A Int B) Int C = A Int (B Int C)
+\idx{Int_Un_distrib} (A Un B) Int C = (A Int C) Un (B Int C)
+
+\idx{Un_absorb} A Un A = A
+\idx{Un_commute} A Un B = B Un A
+\idx{Un_assoc} (A Un B) Un C = A Un (B Un C)
+\idx{Un_Int_distrib} (A Int B) Un C = (A Un C) Int (B Un C)
+
+\idx{Compl_disjoint} A Int Compl(A) = \{x.False\}
+\idx{Compl_partition} A Un Compl(A) = \{x.True\}
+\idx{double_complement} Compl(Compl(A)) = A
+\idx{Compl_Un} Compl(A Un B) = Compl(A) Int Compl(B)
+\idx{Compl_Int} Compl(A Int B) = Compl(A) Un Compl(B)
+
+\idx{Union_Un_distrib} Union(A Un B) = Union(A) Un Union(B)
+\idx{Int_Union_image} A Int Union(B) = (UN C:B. A Int C)
+\idx{Un_Union_image}
+ (UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)
+
+\idx{Inter_Un_distrib} Inter(A Un B) = Inter(A) Int Inter(B)
+\idx{Un_Inter_image} A Un Inter(B) = (INT C:B. A Un C)
+\idx{Int_Inter_image}
+ (INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)
+\end{ttbox}
+\caption{Set equalities} \label{hol-equalities}
+\end{figure}
+
+
+\subsection{Derived rules for sets}
+Figures~\ref{hol-set1} and~\ref{hol-set2} present derived rules. Most
+are obvious and resemble rules of Isabelle's {\ZF} set theory. The
+rules named $XXX${\tt_cong} break down equalities. Certain rules, such as
+\ttindexbold{subsetCE}, \ttindexbold{bexCI} and \ttindexbold{UnCI}, are
+designed for classical reasoning; the more natural rules \ttindexbold{subsetD},
+\ttindexbold{bexI}, \ttindexbold{Un1} and~\ttindexbold{Un2} are not
+strictly necessary. Similarly, \ttindexbold{equalityCE} supports classical
+reasoning about extensionality, after the fashion of \ttindex{iffCE}. See
+the file \ttindexbold{HOL/set.ML} for proofs pertaining to set theory.
+
+Figure~\ref{hol-fun} presents derived rules involving functions. See
+the file \ttindexbold{HOL/fun.ML} for a complete listing.
+
+Figure~\ref{hol-subset} presents lattice properties of the subset relation.
+See \ttindexbold{HOL/subset.ML}.
+
+Figure~\ref{hol-equalities} presents set equalities. See
+\ttindexbold{HOL/equalities.ML}.
+
+
+\begin{figure} \makeatother
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{Pair} & $[\alpha,\beta]\To \alpha\times\beta$
+ & ordered pairs $\langle a,b\rangle$ \\
+ \idx{fst} & $\alpha\times\beta \To \alpha$ & first projection\\
+ \idx{snd} & $\alpha\times\beta \To \beta$ & second projection\\
+ \idx{split} & $[\alpha\times\beta, [\alpha,\beta]\To\gamma] \To \gamma$
+ & generalized projection
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{ttbox}
+\idx{fst_def} fst(p) == @a. ? b. p = <a,b>
+\idx{snd_def} snd(p) == @b. ? a. p = <a,b>
+\idx{split_def} split(p,c) == c(fst(p),snd(p))
+\subcaption{Definitions}
+
+\idx{Pair_inject} [| <a, b> = <a',b'>; [| a=a'; b=b' |] ==> R |] ==> R
+
+\idx{fst} fst(<a,b>) = a
+\idx{snd} snd(<a,b>) = b
+\idx{split} split(<a,b>, c) = c(a,b)
+
+\idx{surjective_pairing} p = <fst(p),snd(p)>
+\subcaption{Derived rules}
+\end{ttbox}
+\caption{Type $\alpha\times\beta$}
+\label{hol-prod}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{Inl} & $\alpha \To \alpha+\beta$ & first injection\\
+ \idx{Inr} & $\beta \To \alpha+\beta$ & second injection\\
+ \idx{case} & $[\alpha+\beta, \alpha\To\gamma, \beta\To\gamma] \To\gamma$
+ & conditional
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{ttbox}
+\idx{case_def} case == (%p f g. @z. (!x. p=Inl(x) --> z=f(x)) &
+ (!y. p=Inr(y) --> z=g(y)))
+\subcaption{Definition}
+
+\idx{Inl_not_Inr} ~ Inl(a)=Inr(b)
+
+\idx{inj_Inl} inj(Inl)
+\idx{inj_Inr} inj(Inr)
+
+\idx{sumE} [| !!x::'a. P(Inl(x)); !!y::'b. P(Inr(y)) |] ==> P(s)
+
+\idx{case_Inl} case(Inl(x), f, g) = f(x)
+\idx{case_Inr} case(Inr(x), f, g) = g(x)
+
+\idx{surjective_sum} case(s, %x::'a. f(Inl(x)), %y::'b. f(Inr(y))) = f(s)
+\subcaption{Derived rules}
+\end{ttbox}
+\caption{Rules for type $\alpha+\beta$}
+\label{hol-sum}
+\end{figure}
+
+
+\section{Types}
+The basic higher-order logic is augmented with a tremendous amount of
+material, including support for recursive function and type definitions.
+Space does not permit a detailed discussion. The present section describes
+product, sum, natural number and list types.
+
+\subsection{Product and sum types}
+{\HOL} defines the product type $\alpha\times\beta$ and the sum type
+$\alpha+\beta$, with the ordered pair syntax {\tt<$a$,$b$>}, using fairly
+standard constructions (Figures~\ref{hol-prod} and~\ref{hol-sum}). Because
+Isabelle does not support type definitions, the isomorphisms between these
+types and their representations are made explicitly.
+
+Most of the definitions are suppressed, but observe that the projections
+and conditionals are defined as descriptions. Their properties are easily
+proved using \ttindex{select_equality}. See \ttindexbold{HOL/prod.thy} and
+\ttindexbold{HOL/sum.thy} for details.
+
+\begin{figure} \makeatother
+\indexbold{*"<}
+\begin{center}
+\begin{tabular}{rrr}
+ \it symbol & \it meta-type & \it description \\
+ \idx{0} & $nat$ & zero \\
+ \idx{Suc} & $nat \To nat$ & successor function\\
+ \idx{nat_case} & $[nat, \alpha, nat\To\alpha] \To\alpha$
+ & conditional\\
+ \idx{nat_rec} & $[nat, \alpha, [nat, \alpha]\To\alpha] \To \alpha$
+ & primitive recursor\\
+ \idx{pred_nat} & $(nat\times nat) set$ & predecessor relation
+\end{tabular}
+\end{center}
+
+\begin{center}
+\indexbold{*"+}
+\index{*@{\tt*}|bold}
+\index{/@{\tt/}|bold}
+\index{//@{\tt//}|bold}
+\index{+@{\tt+}|bold}
+\index{-@{\tt-}|bold}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt * & $[nat,nat]\To nat$ & Left 70 & multiplication \\
+ \tt / & $[nat,nat]\To nat$ & Left 70 & division\\
+ \tt // & $[nat,nat]\To nat$ & Left 70 & modulus\\
+ \tt + & $[nat,nat]\To nat$ & Left 65 & addition\\
+ \tt - & $[nat,nat]\To nat$ & Left 65 & subtraction
+\end{tabular}
+\end{center}
+\subcaption{Constants and infixes}
+
+\begin{ttbox}
+\idx{nat_case_def} nat_case == (%n a f. @z. (n=0 --> z=a) &
+ (!x. n=Suc(x) --> z=f(x)))
+\idx{pred_nat_def} pred_nat == \{p. ? n. p = <n, Suc(n)>\}
+\idx{less_def} m<n == <m,n>:pred_nat^+
+\idx{nat_rec_def} nat_rec(n,c,d) ==
+ wfrec(pred_nat, n, %l g.nat_case(l, c, %m.d(m,g(m))))
+
+\idx{add_def} m+n == nat_rec(m, n, %u v.Suc(v))
+\idx{diff_def} m-n == nat_rec(n, m, %u v. nat_rec(v, 0, %x y.x))
+\idx{mult_def} m*n == nat_rec(m, 0, %u v. n + v)
+\idx{mod_def} m//n == wfrec(trancl(pred_nat), m, %j f. if(j<n,j,f(j-n)))
+\idx{quo_def} m/n == wfrec(trancl(pred_nat),
+ m, %j f. if(j<n,0,Suc(f(j-n))))
+\subcaption{Definitions}
+\end{ttbox}
+\caption{Defining $nat$, the type of natural numbers} \label{hol-nat1}
+\end{figure}
+
+
+\begin{figure} \makeatother
+\begin{ttbox}
+\idx{nat_induct} [| P(0); !!k. [| P(k) |] ==> P(Suc(k)) |] ==> P(n)
+
+\idx{Suc_not_Zero} Suc(m) ~= 0
+\idx{inj_Suc} inj(Suc)
+\idx{n_not_Suc_n} n~=Suc(n)
+\subcaption{Basic properties}
+
+\idx{pred_natI} <n, Suc(n)> : pred_nat
+\idx{pred_natE}
+ [| p : pred_nat; !!x n. [| p = <n, Suc(n)> |] ==> R |] ==> R
+
+\idx{nat_case_0} nat_case(0, a, f) = a
+\idx{nat_case_Suc} nat_case(Suc(k), a, f) = f(k)
+
+\idx{wf_pred_nat} wf(pred_nat)
+\idx{nat_rec_0} nat_rec(0,c,h) = c
+\idx{nat_rec_Suc} nat_rec(Suc(n), c, h) = h(n, nat_rec(n,c,h))
+\subcaption{Case analysis and primitive recursion}
+
+\idx{less_trans} [| i<j; j<k |] ==> i<k
+\idx{lessI} n < Suc(n)
+\idx{zero_less_Suc} 0 < Suc(n)
+
+\idx{less_not_sym} n<m --> ~ m<n
+\idx{less_not_refl} ~ n<n
+\idx{not_less0} ~ n<0
+
+\idx{Suc_less_eq} (Suc(m) < Suc(n)) = (m<n)
+\idx{less_induct} [| !!n. [| ! m. m<n --> P(m) |] ==> P(n) |] ==> P(n)
+
+\idx{less_linear} m<n | m=n | n<m
+\subcaption{The less-than relation}
+\end{ttbox}
+\caption{Derived rules for~$nat$} \label{hol-nat2}
+\end{figure}
+
+
+\subsection{The type of natural numbers, $nat$}
+{\HOL} defines the natural numbers in a roundabout but traditional way.
+The axiom of infinity postulates an type~$ind$ of individuals, which is
+non-empty and closed under an injective operation. The natural numbers are
+inductively generated by choosing an arbitrary individual for~0 and using
+the injective operation to take successors. As usual, the isomorphisms
+between~$nat$ and its representation are made explicitly.
+
+The definition makes use of a least fixed point operator \ttindex{lfp},
+defined using the Knaster-Tarski theorem. This in turn defines an operator
+\ttindex{trancl} for taking the transitive closure of a relation. See
+files \ttindexbold{HOL/lfp.thy} and \ttindexbold{HOL/trancl.thy} for
+details. The definition of~$nat$ resides on \ttindexbold{HOL/nat.thy}.
+
+Type $nat$ is postulated to belong to class~$ord$, which overloads $<$ and
+$\leq$ on the natural numbers. As of this writing, Isabelle provides no
+means of verifying that such overloading is sensible. On the other hand,
+the {\HOL} theory includes no polymorphic axioms stating general properties
+of $<$ and $\leq$.
+
+File \ttindexbold{HOL/arith.ML} develops arithmetic on the natural numbers.
+It defines addition, multiplication, subtraction, division, and remainder,
+proving the theorem $a \bmod b + (a/b)\times b = a$. Division and
+remainder are defined by repeated subtraction, which requires well-founded
+rather than primitive recursion.
+
+Primitive recursion makes use of \ttindex{wfrec}, an operator for recursion
+along arbitrary well-founded relations; see \ttindexbold{HOL/wf.ML} for the
+development. The predecessor relation, \ttindexbold{pred_nat}, is shown to
+be well-founded; recursion along this relation is primitive recursion,
+while its transitive closure is~$<$.
+
+
+\begin{figure} \makeatother
+\begin{center}
+\begin{tabular}{rrr}
+ \it symbol & \it meta-type & \it description \\
+ \idx{Nil} & $\alpha list$ & the empty list\\
+ \idx{Cons} & $[\alpha, \alpha list] \To \alpha list$
+ & list constructor\\
+ \idx{list_rec} & $[\alpha list, \beta, [\alpha ,\alpha list,
+\beta]\To\beta] \To \beta$
+ & list recursor\\
+ \idx{map} & $(\alpha\To\beta) \To (\alpha list \To \beta list)$
+ & mapping functional
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{ttbox}
+\idx{map_def} map(f,xs) == list_rec(xs, Nil, %x l r. Cons(f(x), r))
+\subcaption{Definition}
+
+\idx{list_induct}
+ [| P(Nil); !!x xs. [| P(xs) |] ==> P(Cons(x,xs)) |] ==> P(l)
+
+\idx{Cons_not_Nil} ~ Cons(x,xs) = Nil
+\idx{Cons_Cons_eq} (Cons(x,xs)=Cons(y,ys)) = (x=y & xs=ys)
+
+\idx{list_rec_Nil} list_rec(Nil,c,h) = c
+\idx{list_rec_Cons} list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))
+
+\idx{map_Nil} map(f,Nil) = Nil
+\idx{map_Cons} map(f, Cons(x,xs)) = Cons(f(x), map(f,xs))
+\end{ttbox}
+\caption{The type of lists and its operations} \label{hol-list}
+\end{figure}
+
+
+\subsection{The type constructor for lists, $\alpha\,list$}
+{\HOL}'s definition of lists is an example of an experimental method for
+handling recursive data types. The details need not concern us here; see
+the file \ttindexbold{HOL/list.ML}. Figure~\ref{hol-list} presents the
+basic list operations, with their types and properties. In particular,
+\ttindexbold{list_rec} is a primitive recursion operator for lists, in the
+style of Martin-L\"of type theory. It is derived from well-founded
+recursion, a general principle that can express arbitrary total recursive
+functions.
+
+
+\subsection{The type constructor for lazy lists, $\alpha\,llist$}
+The definition of lazy lists demonstrates methods for handling infinite
+data structures and co-induction in higher-order logic. It defines an
+operator for co-recursion on lazy lists, which is used to define a few
+simple functions such as map and append. Co-recursion cannot easily define
+operations such as filter, which can compute indefinitely before yielding
+the next element (if any!) of the lazy list. A co-induction principle is
+defined for proving equations on lazy lists. See the files
+\ttindexbold{HOL/llist.thy} and \ttindexbold{HOL/llist.ML} for the formal
+derivations. I have written a report discussing the treatment of lazy
+lists, and finite lists also~\cite{paulson-coind}.
+
+
+\section{Classical proof procedures} \label{hol-cla-prover}
+{\HOL} derives classical introduction rules for $\disj$ and~$\exists$, as
+well as classical elimination rules for~$\imp$ and~$\bimp$, and the swap
+rule (Figure~\ref{hol-lemmas2}).
+
+The classical reasoning module is set up for \HOL, as the structure
+\ttindexbold{Classical}. This structure is open, so {\ML} identifiers such
+as {\tt step_tac}, {\tt fast_tac}, {\tt best_tac}, etc., refer to it.
+
+{\HOL} defines the following classical rule sets:
+\begin{ttbox}
+prop_cs : claset
+HOL_cs : claset
+HOL_dup_cs : claset
+set_cs : claset
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{prop_cs}] contains the propositional rules, namely
+those for~$\top$, $\bot$, $\conj$, $\disj$, $\neg$, $\imp$ and~$\bimp$,
+along with the rule~\ttindex{refl}.
+
+\item[\ttindexbold{HOL_cs}]
+extends {\tt prop_cs} with the safe rules \ttindex{allI} and~\ttindex{exE}
+and the unsafe rules \ttindex{allE} and~\ttindex{exI}, as well as rules for
+unique existence. Search using this is incomplete since quantified
+formulae are used at most once.
+
+\item[\ttindexbold{HOL_dup_cs}]
+extends {\tt prop_cs} with the safe rules \ttindex{allI} and~\ttindex{exE}
+and the unsafe rules \ttindex{all_dupE} and~\ttindex{exCI}, as well as
+rules for unique existence. Search using this is complete --- quantified
+formulae may be duplicated --- but frequently fails to terminate. It is
+generally unsuitable for depth-first search.
+
+\item[\ttindexbold{set_cs}]
+extends {\tt HOL_cs} with rules for the bounded quantifiers, subsets,
+comprehensions, unions/intersections, complements, finite setes, images and
+ranges.
+\end{description}
+\noindent
+See the {\em Reference Manual} for more discussion of classical proof
+methods.
+
+
+\section{The examples directory}
+This directory contains examples and experimental proofs in {\HOL}. Here
+is an overview of the more interesting files.
+\begin{description}
+\item[\ttindexbold{HOL/ex/meson.ML}]
+contains an experimental implementation of the MESON proof procedure,
+inspired by Plaisted~\cite{plaisted90}. It is much more powerful than
+Isabelle's classical module.
+
+\item[\ttindexbold{HOL/ex/meson-test.ML}]
+contains test data for the MESON proof procedure.
+
+\item[\ttindexbold{HOL/ex/set.ML}]
+proves Cantor's Theorem (see below) and the Schr\"oder-Bernstein Theorem.
+
+\item[\ttindexbold{HOL/ex/prop-log.ML}]
+proves the soundness and completeness of classical propositional logic,
+given a truth table semantics. The only connective is $\imp$. A
+Hilbert-style axiom system is specified, and its set of theorems defined
+inductively.
+
+\item[\ttindexbold{HOL/ex/term.ML}]
+is an experimental recursive type definition, where the recursion goes
+through the type constructor~$list$.
+
+\item[\ttindexbold{HOL/ex/simult.ML}]
+defines primitives for solving mutually recursive equations over sets.
+It constructs sets of trees and forests as an example, including induction
+and recursion rules that handle the mutual recursion.
+\end{description}
+
+
+\section{Example: deriving the conjunction rules}
+{\HOL} comes with a body of derived rules, ranging from simple properties
+of the logical constants and set theory to well-founded recursion. Many of
+them are worth studying.
+
+Deriving natural deduction rules for the logical constants from their
+definitions is an archetypal example of higher-order reasoning. Let us
+verify two conjunction rules:
+\[ \infer[({\conj}I)]{P\conj Q}{P & Q} \qquad\qquad
+ \infer[({\conj}E1)]{P}{P\conj Q}
+\]
+
+\subsection{The introduction rule}
+We begin by stating the rule as the goal. The list of premises $[P,Q]$ is
+bound to the {\ML} variable~{\tt prems}.
+\begin{ttbox}
+val prems = goal HOL.thy "[| P; Q |] ==> P&Q";
+{\out Level 0}
+{\out P & Q}
+{\out 1. P & Q}
+\end{ttbox}
+The next step is to unfold the definition of conjunction. But
+\ttindex{and_def} uses {\HOL}'s internal equality, so
+\ttindex{rewrite_goals_tac} is unsuitable.
+Instead, we perform substitution using the rule \ttindex{ssubst}:
+\begin{ttbox}
+by (resolve_tac [and_def RS ssubst] 1);
+{\out Level 1}
+{\out P & Q}
+{\out 1. ! R. (P --> Q --> R) --> R}
+\end{ttbox}
+We now apply $(\forall I)$ and $({\imp}I)$:
+\begin{ttbox}
+by (resolve_tac [allI] 1);
+{\out Level 2}
+{\out P & Q}
+{\out 1. !!R. (P --> Q --> R) --> R}
+by (resolve_tac [impI] 1);
+{\out Level 3}
+{\out P & Q}
+{\out 1. !!R. P --> Q --> R ==> R}
+\end{ttbox}
+The assumption is a nested implication, which may be eliminated
+using~\ttindex{mp} resolved with itself. Elim-resolution, here, performs
+backwards chaining. More straightforward would be to use~\ttindex{impE}
+twice.
+\index{*RS}
+\begin{ttbox}
+by (eresolve_tac [mp RS mp] 1);
+{\out Level 4}
+{\out P & Q}
+{\out 1. !!R. P}
+{\out 2. !!R. Q}
+\end{ttbox}
+These two subgoals are simply the premises:
+\begin{ttbox}
+by (REPEAT (resolve_tac prems 1));
+{\out Level 5}
+{\out P & Q}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\subsection{The elimination rule}
+Again, we bind the list of premises (in this case $[P\conj Q]$)
+to~{\tt prems}.
+\begin{ttbox}
+val prems = goal HOL.thy "[| P & Q |] ==> P";
+{\out Level 0}
+{\out P}
+{\out 1. P}
+\end{ttbox}
+Working with premises that involve defined constants can be tricky. We
+must expand the definition of conjunction in the meta-assumption $P\conj
+Q$. The rule \ttindex{subst} performs substitution in forward proofs.
+We get two resolvents, since the vacuous substitution is valid:
+\begin{ttbox}
+prems RL [and_def RS subst];
+{\out val it = ["! R. (P --> Q --> R) --> R [P & Q]",}
+{\out "P & Q [P & Q]"] : thm list}
+\end{ttbox}
+By applying $(\forall E)$ and $({\imp}E)$ to the resolvents, we dispose of
+the vacuous one and put the other into a convenient form:
+\index{*RL}
+\begin{ttbox}
+prems RL [and_def RS subst] RL [spec] RL [mp];
+{\out val it = ["P --> Q --> ?Q ==> ?Q [P & Q]"] : thm list}
+\end{ttbox}
+This is a list containing a single rule, which is directly applicable to
+our goal:
+\begin{ttbox}
+by (resolve_tac it 1);
+{\out Level 1}
+{\out P}
+{\out 1. P --> Q --> P}
+\end{ttbox}
+The subgoal is a trivial implication. Recall that \ttindex{ares_tac} is a
+combination of \ttindex{assume_tac} and \ttindex{resolve_tac}.
+\begin{ttbox}
+by (REPEAT (ares_tac [impI] 1));
+{\out Level 2}
+{\out P}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\section{Example: Cantor's Theorem}
+Cantor's Theorem states that every set has more subsets than it has
+elements. It has become a favourite example in higher-order logic since
+it is so easily expressed:
+\[ \forall f::[\alpha,\alpha]\To bool. \exists S::\alpha\To bool.
+ \forall x::\alpha. f(x) \not= S
+\]
+Viewing types as sets, $\alpha\To bool$ represents the powerset
+of~$\alpha$. This version states that for every function from $\alpha$ to
+its powerset, some subset is outside its range.
+
+The Isabelle proof uses {\HOL}'s set theory, with the type $\alpha\,set$ and
+the operator \ttindex{range}. Since it avoids quantification, we may
+inspect the subset found by the proof.
+\begin{ttbox}
+goal Set.thy "~ ?S : range(f :: 'a=>'a set)";
+{\out Level 0}
+{\out ~ ?S : range(f)}
+{\out 1. ~ ?S : range(f)}
+\end{ttbox}
+The first two steps are routine. The rule \ttindex{rangeE} reasons that,
+since $\Var{S}\in range(f)$, we have $\Var{S}=f(x)$ for some~$x$.
+\begin{ttbox}
+by (resolve_tac [notI] 1);
+{\out Level 1}
+{\out ~ ?S : range(f)}
+{\out 1. ?S : range(f) ==> False}
+by (eresolve_tac [rangeE] 1);
+{\out Level 2}
+{\out ~ ?S : range(f)}
+{\out 1. !!x. ?S = f(x) ==> False}
+\end{ttbox}
+Next, we apply \ttindex{equalityCE}, reasoning that since $\Var{S}=f(x)$,
+we have $\Var{c}\in \Var{S}$ if and only if $\Var{c}\in f(x)$ for
+any~$\Var{c}$.
+\begin{ttbox}
+by (eresolve_tac [equalityCE] 1);
+{\out Level 3}
+{\out ~ ?S : range(f)}
+{\out 1. !!x. [| ?c3(x) : ?S; ?c3(x) : f(x) |] ==> False}
+{\out 2. !!x. [| ~ ?c3(x) : ?S; ~ ?c3(x) : f(x) |] ==> False}
+\end{ttbox}
+Now we use a bit of creativity. Suppose that $\Var{S}$ has the form of a
+comprehension. Then $\Var{c}\in\{x.\Var{P}(x)\}$ implies
+$\Var{P}(\Var{c})\}$.\index{*CollectD}
+\begin{ttbox}
+by (dresolve_tac [CollectD] 1);
+{\out Level 4}
+{\out ~ \{x. ?P7(x)\} : range(f)}
+{\out 1. !!x. [| ?c3(x) : f(x); ?P7(?c3(x)) |] ==> False}
+{\out 2. !!x. [| ~ ?c3(x) : \{x. ?P7(x)\}; ~ ?c3(x) : f(x) |] ==> False}
+\end{ttbox}
+Forcing a contradiction between the two assumptions of subgoal~1 completes
+the instantiation of~$S$. It is now $\{x. x\not\in f(x)\}$, the standard
+diagonal construction.
+\begin{ttbox}
+by (contr_tac 1);
+{\out Level 5}
+{\out ~ \{x. ~ x : f(x)\} : range(f)}
+{\out 1. !!x. [| ~ x : \{x. ~ x : f(x)\}; ~ x : f(x) |] ==> False}
+\end{ttbox}
+The rest should be easy. To apply \ttindex{CollectI} to the negated
+assumption, we employ \ttindex{swap_res_tac}:
+\begin{ttbox}
+by (swap_res_tac [CollectI] 1);
+{\out Level 6}
+{\out ~ \{x. ~ x : f(x)\} : range(f)}
+{\out 1. !!x. [| ~ x : f(x); ~ False |] ==> ~ x : f(x)}
+by (assume_tac 1);
+{\out Level 7}
+{\out ~ \{x. ~ x : f(x)\} : range(f)}
+{\out No subgoals!}
+\end{ttbox}
+How much creativity is required? As it happens, Isabelle can prove this
+theorem automatically. The classical set \ttindex{set_cs} contains rules
+for most of the constructs of {\HOL}'s set theory. We augment it with
+\ttindex{equalityCE} --- set equalities are not broken up by default ---
+and apply best-first search. Depth-first search would diverge, but
+best-first search successfully navigates through the large search space.
+\begin{ttbox}
+choplev 0;
+{\out Level 0}
+{\out ~ ?S : range(f)}
+{\out 1. ~ ?S : range(f)}
+by (best_tac (set_cs addSEs [equalityCE]) 1);
+{\out Level 1}
+{\out ~ \{x. ~ x : f(x)\} : range(f)}
+{\out No subgoals!}
+\end{ttbox}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/ZF-eg.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,217 @@
+(**** ZF examples ****)
+
+Pretty.setmargin 72; (*existing macros just allow this margin*)
+print_depth 0;
+
+(*** Powerset example ***)
+
+val [prem] = goal ZF_Rule.thy "A<=B ==> Pow(A) <= Pow(B)";
+by (resolve_tac [subsetI] 1);
+by (resolve_tac [PowI] 1);
+by (dresolve_tac [PowD] 1);
+by (eresolve_tac [subset_trans] 1);
+by (resolve_tac [prem] 1);
+val Pow_mono = result();
+
+goal ZF_Rule.thy "Pow(A Int B) = Pow(A) Int Pow(B)";
+by (resolve_tac [equalityI] 1);
+by (resolve_tac [Int_greatest] 1);
+by (resolve_tac [Int_lower1 RS Pow_mono] 1);
+by (resolve_tac [Int_lower2 RS Pow_mono] 1);
+by (resolve_tac [subsetI] 1);
+by (eresolve_tac [IntE] 1);
+by (resolve_tac [PowI] 1);
+by (REPEAT (dresolve_tac [PowD] 1));
+by (resolve_tac [Int_greatest] 1);
+by (REPEAT (assume_tac 1));
+choplev 0;
+by (fast_tac (ZF_cs addIs [equalityI]) 1);
+
+val [prem] = goal ZF_Rule.thy "C<=D ==> Union(C) <= Union(D)";
+by (resolve_tac [subsetI] 1);
+by (eresolve_tac [UnionE] 1);
+by (resolve_tac [UnionI] 1);
+by (resolve_tac [prem RS subsetD] 1);
+by (assume_tac 1);
+by (assume_tac 1);
+choplev 0;
+by (resolve_tac [Union_least] 1);
+by (resolve_tac [Union_upper] 1);
+by (eresolve_tac [prem RS subsetD] 1);
+
+
+val prems = goal ZF_Rule.thy
+ "[| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \
+\ (f Un g)`a = f`a";
+by (resolve_tac [apply_equality] 1);
+by (resolve_tac [UnI1] 1);
+by (resolve_tac [apply_Pair] 1);
+by (resolve_tac prems 1);
+by (resolve_tac prems 1);
+by (resolve_tac [fun_disjoint_Un] 1);
+by (resolve_tac prems 1);
+by (resolve_tac prems 1);
+by (resolve_tac prems 1);
+
+
+
+
+goal ZF_Rule.thy "f``(UN x:A. B(x)) = (UN x:A. f``B(x))";
+by (resolve_tac [equalityI] 1);
+by (resolve_tac [subsetI] 1);
+fe imageE;
+
+
+goal ZF_Rule.thy "(UN x:C. A(x) Int B) = (UN x:C. A(x)) Int B";
+by (resolve_tac [equalityI] 1);
+by (resolve_tac [Int_greatest] 1);
+fr UN_mono;
+by (resolve_tac [Int_lower1] 1);
+fr UN_least;
+????
+
+
+> goal ZF_Rule.thy "Pow(A Int B) = Pow(A) Int Pow(B)";
+Level 0
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A Int B) = Pow(A) Int Pow(B)
+> by (resolve_tac [equalityI] 1);
+Level 1
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A Int B) <= Pow(A) Int Pow(B)
+ 2. Pow(A) Int Pow(B) <= Pow(A Int B)
+> by (resolve_tac [Int_greatest] 1);
+Level 2
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A Int B) <= Pow(A)
+ 2. Pow(A Int B) <= Pow(B)
+ 3. Pow(A) Int Pow(B) <= Pow(A Int B)
+> by (resolve_tac [Int_lower1 RS Pow_mono] 1);
+Level 3
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A Int B) <= Pow(B)
+ 2. Pow(A) Int Pow(B) <= Pow(A Int B)
+> by (resolve_tac [Int_lower2 RS Pow_mono] 1);
+Level 4
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A) Int Pow(B) <= Pow(A Int B)
+> by (resolve_tac [subsetI] 1);
+Level 5
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. !!x. x : Pow(A) Int Pow(B) ==> x : Pow(A Int B)
+> by (eresolve_tac [IntE] 1);
+Level 6
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. !!x. [| x : Pow(A); x : Pow(B) |] ==> x : Pow(A Int B)
+> by (resolve_tac [PowI] 1);
+Level 7
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. !!x. [| x : Pow(A); x : Pow(B) |] ==> x <= A Int B
+> by (REPEAT (dresolve_tac [PowD] 1));
+Level 8
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. !!x. [| x <= A; x <= B |] ==> x <= A Int B
+> by (resolve_tac [Int_greatest] 1);
+Level 9
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. !!x. [| x <= A; x <= B |] ==> x <= A
+ 2. !!x. [| x <= A; x <= B |] ==> x <= B
+> by (REPEAT (assume_tac 1));
+Level 10
+Pow(A Int B) = Pow(A) Int Pow(B)
+No subgoals!
+> choplev 0;
+Level 0
+Pow(A Int B) = Pow(A) Int Pow(B)
+ 1. Pow(A Int B) = Pow(A) Int Pow(B)
+> by (fast_tac (ZF_cs addIs [equalityI]) 1);
+Level 1
+Pow(A Int B) = Pow(A) Int Pow(B)
+No subgoals!
+
+
+
+
+> val [prem] = goal ZF_Rule.thy "C<=D ==> Union(C) <= Union(D)";
+Level 0
+Union(C) <= Union(D)
+ 1. Union(C) <= Union(D)
+> by (resolve_tac [subsetI] 1);
+Level 1
+Union(C) <= Union(D)
+ 1. !!x. x : Union(C) ==> x : Union(D)
+> by (eresolve_tac [UnionE] 1);
+Level 2
+Union(C) <= Union(D)
+ 1. !!x B. [| x : B; B : C |] ==> x : Union(D)
+> by (resolve_tac [UnionI] 1);
+Level 3
+Union(C) <= Union(D)
+ 1. !!x B. [| x : B; B : C |] ==> ?B2(x,B) : D
+ 2. !!x B. [| x : B; B : C |] ==> x : ?B2(x,B)
+> by (resolve_tac [prem RS subsetD] 1);
+Level 4
+Union(C) <= Union(D)
+ 1. !!x B. [| x : B; B : C |] ==> ?B2(x,B) : C
+ 2. !!x B. [| x : B; B : C |] ==> x : ?B2(x,B)
+> by (assume_tac 1);
+Level 5
+Union(C) <= Union(D)
+ 1. !!x B. [| x : B; B : C |] ==> x : B
+> by (assume_tac 1);
+Level 6
+Union(C) <= Union(D)
+No subgoals!
+
+
+
+> val prems = goal ZF_Rule.thy
+# "[| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \
+# \ (f Un g)`a = f`a";
+Level 0
+(f Un g) ` a = f ` a
+ 1. (f Un g) ` a = f ` a
+> by (resolve_tac [apply_equality] 1);
+Level 1
+(f Un g) ` a = f ` a
+ 1. <a,f ` a> : f Un g
+ 2. f Un g : (PROD x:?A. ?B(x))
+> by (resolve_tac [UnI1] 1);
+Level 2
+(f Un g) ` a = f ` a
+ 1. <a,f ` a> : f
+ 2. f Un g : (PROD x:?A. ?B(x))
+> by (resolve_tac [apply_Pair] 1);
+Level 3
+(f Un g) ` a = f ` a
+ 1. f : (PROD x:?A2. ?B2(x))
+ 2. a : ?A2
+ 3. f Un g : (PROD x:?A. ?B(x))
+> by (resolve_tac prems 1);
+Level 4
+(f Un g) ` a = f ` a
+ 1. a : A
+ 2. f Un g : (PROD x:?A. ?B(x))
+> by (resolve_tac prems 1);
+Level 5
+(f Un g) ` a = f ` a
+ 1. f Un g : (PROD x:?A. ?B(x))
+> by (resolve_tac [fun_disjoint_Un] 1);
+Level 6
+(f Un g) ` a = f ` a
+ 1. f : ?A3 -> ?B3
+ 2. g : ?C3 -> ?D3
+ 3. ?A3 Int ?C3 = 0
+> by (resolve_tac prems 1);
+Level 7
+(f Un g) ` a = f ` a
+ 1. g : ?C3 -> ?D3
+ 2. A Int ?C3 = 0
+> by (resolve_tac prems 1);
+Level 8
+(f Un g) ` a = f ` a
+ 1. A Int C = 0
+> by (resolve_tac prems 1);
+Level 9
+(f Un g) ` a = f ` a
+No subgoals!
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/ZF-rules.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,468 @@
+%%%% RULES.ML
+
+\idx{empty_set} ~(x:0)
+\idx{union_iff} A:Union(C) <-> (EX B:C. A:B)
+\idx{power_set} A : Pow(B) <-> A <= B
+\idx{infinity} 0:Inf & (ALL y:Inf. succ(y): Inf)
+\idx{foundation} A=0 | (EX x:A. ALL y:x. ~ y:A)
+
+\idx{replacement} (!!x y z.[| x:A; P(x,y); P(x,z) |] ==> y=z) ==>
+ y : PrimReplace(A,P) <-> (EX x:A. P(x,y))
+
+\idx{Replace_def} Replace(A,P) == PrimReplace(A, %x y. (EX!z.P(x,z)) & P(x,y))
+\idx{RepFun_def} RepFun(A,f) == Replace(A, %x u. u=f(x))
+\idx{Collect_def} Collect(A,P) == \{ y . x:A, x=y & P(x)\}
+\idx{the_def} The(P) == Union(\{y . x:\{0\}, P(y)\})
+
+\idx{Upair_def} Upair(a,b) ==
+ \{y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)\}
+
+\idx{Inter_def} Inter(A) == \{ x:Union(A) . ALL y:A. x:y\}
+
+\idx{Un_def} A Un B == Union(Upair(A,B))
+\idx{Int_def} A Int B == Inter(Upair(A,B))
+\idx{Diff_def} A - B == \{ x:A . ~(x:B) \}
+\idx{cons_def} cons(a,A) == Upair(a,a) Un A
+\idx{succ_def} succ(i) == cons(i,i)
+
+\idx{Pair_def} <a,b> == \{\{a,a\}, \{a,b\}\}
+\idx{fst_def} fst(A) == THE x. EX y. A=<x,y>
+\idx{snd_def} snd(A) == THE y. EX x. A=<x,y>
+\idx{split_def} split(p,c) == THE y. EX a b. p=<a,b> & y=c(a,b)
+\idx{Sigma_def} Sigma(A,B) == UN x:A. UN y:B(x). \{<x,y>\}
+
+\idx{domain_def} domain(r) == \{a:Union(Union(r)) . EX b. <a,b> : r\}
+\idx{range_def} range(r) == \{b:Union(Union(r)) . EX a. <a,b> : r\}
+\idx{field_def} field(r) == domain(r) Un range(r)
+\idx{image_def} r``A == \{y : range(r) . EX x:A. <x,y> : r\}
+\idx{vimage_def} r -`` A == \{x : domain(r) . EX y:A. <x,y> : r\}
+
+\idx{lam_def} Lambda(A,f) == RepFun(A, %x. <x,f(x)>)
+\idx{apply_def} f`a == THE y. <a,y> : f
+\idx{restrict_def} restrict(f,A) == lam x:A.f`x
+\idx{Pi_def} Pi(A,B) == \{f: Pow(Sigma(A,B)). ALL x:A. EX! y. <x,y>: f\}
+
+\idx{subset_def} A <= B == ALL x:A. x:B
+\idx{strict_subset_def} A <! B == A <=B & ~(A=B)
+\idx{extension} A = B <-> A <= B & B <= A
+
+\idx{Ball_def} Ball(A,P) == ALL x. x:A --> P(x)
+\idx{Bex_def} Bex(A,P) == EX x. x:A & P(x)
+
+
+%%%% LEMMAS.ML
+
+\idx{ballI} [| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)
+\idx{bspec} [| ALL x:A. P(x); x: A |] ==> P(x)
+\idx{ballE} [| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q
+
+\idx{ball_cong} [| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==>
+ (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))
+
+\idx{bexI} [| P(x); x: A |] ==> EX x:A. P(x)
+\idx{bexCI} [| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A.P(x)
+\idx{bexE} [| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q
+
+\idx{bex_cong} [| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==>
+ (EX x:A. P(x)) <-> (EX x:A'. P'(x))
+
+\idx{subsetI} (!!x.x:A ==> x:B) ==> A <= B
+\idx{subsetD} [| A <= B; c:A |] ==> c:B
+\idx{subsetCE} [| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P
+\idx{subset_refl} A <= A
+\idx{subset_trans} [| A<=B; B<=C |] ==> A<=C
+
+\idx{equalityI} [| A <= B; B <= A |] ==> A = B
+\idx{equalityD1} A = B ==> A<=B
+\idx{equalityD2} A = B ==> B<=A
+\idx{equalityE} [| A = B; [| A<=B; B<=A |] ==> P |] ==> P
+
+\idx{emptyE} a:0 ==> P
+\idx{empty_subsetI} 0 <= A
+\idx{equals0I} [| !!y. y:A ==> False |] ==> A=0
+\idx{equals0D} [| A=0; a:A |] ==> P
+
+\idx{PowI} A <= B ==> A : Pow(B)
+\idx{PowD} A : Pow(B) ==> A<=B
+
+\idx{ReplaceI} [| x: A; P(x,b); !!y. P(x,y) ==> y=b |] ==>
+ b : \{y. x:A, P(x,y)\}
+
+\idx{ReplaceE} [| b : \{y. x:A, P(x,y)\};
+ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R
+ |] ==> R
+
+\idx{Replace_cong} [| A=B; !!x y. x:B ==> P(x,y) <-> Q(x,y) |] ==>
+ \{y. x:A, P(x,y)\} = \{y. x:B, Q(x,y)\}
+
+\idx{RepFunI} [| a : A |] ==> f(a) : RepFun(A,f)
+\idx{RepFunE} [| b : RepFun(A, %x.f(x));
+ !!x.[| x:A; b=f(x) |] ==> P |] ==> P
+
+\idx{RepFun_cong} [| A=B; !!x. x:B ==> f(x)=g(x) |] ==>
+ RepFun(A, %x.f(x)) = RepFun(B, %x.g(x))
+
+
+\idx{separation} x : Collect(A,P) <-> x:A & P(x)
+\idx{CollectI} [| a:A; P(a) |] ==> a : \{x:A. P(x)\}
+\idx{CollectE} [| a : \{x:A. P(x)\}; [| a:A; P(a) |] ==> R |] ==> R
+\idx{CollectD1} a : \{x:A. P(x)\} ==> a:A
+\idx{CollectD2} a : \{x:A. P(x)\} ==> P(a)
+
+\idx{Collect_cong} [| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==>
+ \{x:A. P(x)\} = \{x:B. Q(x)\}
+
+\idx{UnionI} [| B: C; A: B |] ==> A: Union(C)
+\idx{UnionE} [| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R
+
+\idx{InterI} [| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)
+\idx{InterD} [| A : Inter(C); B : C |] ==> A : B
+\idx{InterE} [| A : Inter(C); A:B ==> R; ~ B:C ==> R |] ==> R
+
+\idx{UN_I} [| a: A; b: B(a) |] ==> b: (UN x:A. B(x))
+\idx{UN_E} [| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R
+
+\idx{INT_I} [| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))
+\idx{INT_E} [| b : (INT x:A. B(x)); a: A |] ==> b : B(a)
+
+
+%%%% UPAIR.ML
+
+\idx{pairing} a:Upair(b,c) <-> (a=b | a=c)
+\idx{UpairI1} a : Upair(a,b)
+\idx{UpairI2} b : Upair(a,b)
+\idx{UpairE} [| a : Upair(b,c); a = b ==> P; a = c ==> P |] ==> P
+
+\idx{UnI1} c : A ==> c : A Un B
+\idx{UnI2} c : B ==> c : A Un B
+\idx{UnCI} (~c : B ==> c : A) ==> c : A Un B
+\idx{UnE} [| c : A Un B; c:A ==> P; c:B ==> P |] ==> P
+
+\idx{IntI} [| c : A; c : B |] ==> c : A Int B
+\idx{IntD1} c : A Int B ==> c : A
+\idx{IntD2} c : A Int B ==> c : B
+\idx{IntE} [| c : A Int B; [| c:A; c:B |] ==> P |] ==> P
+
+\idx{DiffI} [| c : A; ~ c : B |] ==> c : A - B
+\idx{DiffD1} c : A - B ==> c : A
+\idx{DiffD2} [| c : A - B; c : B |] ==> P
+\idx{DiffE} [| c : A - B; [| c:A; ~ c:B |] ==> P |] ==> P
+
+\idx{consI1} a : cons(a,B)
+\idx{consI2} a : B ==> a : cons(b,B)
+\idx{consCI} (~ a:B ==> a=b) ==> a: cons(b,B)
+\idx{consE} [| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P
+
+\idx{singletonI} a : \{a\}
+\idx{singletonE} [| a : \{b\}; a=b ==> P |] ==> P
+
+\idx{succI1} i : succ(i)
+\idx{succI2} i : j ==> i : succ(j)
+\idx{succCI} (~ i:j ==> i=j) ==> i: succ(j)
+\idx{succE} [| i : succ(j); i=j ==> P; i:j ==> P |] ==> P
+\idx{succ_neq_0} [| succ(n)=0 |] ==> P
+\idx{succ_inject} succ(m) = succ(n) ==> m=n
+
+\idx{the_equality} [| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a
+\idx{theI} EX! x. P(x) ==> P(THE x. P(x))
+
+\idx{mem_anti_sym} [| a:b; b:a |] ==> P
+\idx{mem_anti_refl} a:a ==> P
+
+
+%%% SUBSET.ML
+
+\idx{Union_upper} B:A ==> B <= Union(A)
+\idx{Union_least} [| !!x. x:A ==> x<=C |] ==> Union(A) <= C
+
+\idx{Inter_lower} B:A ==> Inter(A) <= B
+\idx{Inter_greatest} [| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)
+
+\idx{Un_upper1} A <= A Un B
+\idx{Un_upper2} B <= A Un B
+\idx{Un_least} [| A<=C; B<=C |] ==> A Un B <= C
+
+\idx{Int_lower1} A Int B <= A
+\idx{Int_lower2} A Int B <= B
+\idx{Int_greatest} [| C<=A; C<=B |] ==> C <= A Int B
+
+\idx{Diff_subset} A-B <= A
+\idx{Diff_contains} [| C<=A; C Int B = 0 |] ==> C <= A-B
+
+\idx{Collect_subset} Collect(A,P) <= A
+
+%%% PAIR.ML
+
+\idx{Pair_inject1} <a,b> = <c,d> ==> a=c
+\idx{Pair_inject2} <a,b> = <c,d> ==> b=d
+\idx{Pair_inject} [| <a,b> = <c,d>; [| a=c; b=d |] ==> P |] ==> P
+\idx{Pair_neq_0} <a,b>=0 ==> P
+
+\idx{fst_conv} fst(<a,b>) = a
+\idx{snd_conv} snd(<a,b>) = b
+\idx{split_conv} split(<a,b>, %x y.c(x,y)) = c(a,b)
+
+\idx{SigmaI} [| a:A; b:B(a) |] ==> <a,b> : (SUM x:A. B(x))
+
+\idx{SigmaE} [| c: (SUM x:A. B(x));
+ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P
+ |] ==> P
+
+\idx{SigmaE2} [| <a,b> : (SUM x:A. B(x));
+ [| a:A; b:B(a) |] ==> P
+ |] ==> P
+
+
+%%% DOMRANGE.ML
+
+\idx{domainI} <a,b>: r ==> a : domain(r)
+\idx{domainE} [| a : domain(r); !!y. <a,y>: r ==> P |] ==> P
+\idx{domain_subset} domain(Sigma(A,B)) <= A
+
+\idx{rangeI} <a,b>: r ==> b : range(r)
+\idx{rangeE} [| b : range(r); !!x. <x,b>: r ==> P |] ==> P
+\idx{range_subset} range(A*B) <= B
+
+\idx{fieldI1} <a,b>: r ==> a : field(r)
+\idx{fieldI2} <a,b>: r ==> b : field(r)
+\idx{fieldCI} (~ <c,a>:r ==> <a,b>: r) ==> a : field(r)
+
+\idx{fieldE} [| a : field(r);
+ !!x. <a,x>: r ==> P;
+ !!x. <x,a>: r ==> P
+ |] ==> P
+
+\idx{field_subset} field(A*A) <= A
+
+\idx{imageI} [| <a,b>: r; a:A |] ==> b : r``A
+\idx{imageE} [| b: r``A; !!x.[| <x,b>: r; x:A |] ==> P |] ==> P
+
+\idx{vimageI} [| <a,b>: r; b:B |] ==> a : r-``B
+\idx{vimageE} [| a: r-``B; !!x.[| <a,x>: r; x:B |] ==> P |] ==> P
+
+
+%%% FUNC.ML
+
+\idx{fun_is_rel} f: (PROD x:A.B(x)) ==> f <= Sigma(A,B)
+
+\idx{apply_equality} [| <a,b>: f; f: (PROD x:A.B(x)) |] ==> f`a = b
+\idx{apply_equality2} [| <a,b>: f; <a,c>: f; f: (PROD x:A.B(x)) |] ==> b=c
+
+\idx{apply_type} [| f: (PROD x:A.B(x)); a:A |] ==> f`a : B(a)
+\idx{apply_Pair} [| f: (PROD x:A.B(x)); a:A |] ==> <a,f`a>: f
+\idx{apply_iff} [| f: (PROD x:A.B(x)); a:A |] ==> <a,b>: f <-> f`a = b
+
+\idx{domain_type} [| <a,b> : f; f: (PROD x:A.B(x)) |] ==> a : A
+\idx{range_type} [| <a,b> : f; f: (PROD x:A.B(x)) |] ==> b : B(a)
+
+\idx{Pi_type} [| f: A->C; !!x. x:A ==> f`x : B(x) |] ==> f: Pi(A,B)
+\idx{domain_of_fun} f : Pi(A,B) ==> domain(f)=A
+\idx{range_of_fun} f : Pi(A,B) ==> f: A->range(f)
+
+\idx{fun_extension} [| f : (PROD x:A.B(x)); g: (PROD x:A.D(x));
+ !!x. x:A ==> f`x = g`x
+ |] ==> f=g
+
+\idx{lamI} a:A ==> <a,b(a)> : (lam x:A. b(x))
+\idx{lamE} [| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P
+ |] ==> P
+
+\idx{lam_type} [| !!x. x:A ==> b(x): B(x) |] ==>
+ (lam x:A.b(x)) : (PROD x:A.B(x))
+
+\idx{beta_conv} a : A ==> (lam x:A.b(x)) ` a = b(a)
+\idx{eta_conv} f : (PROD x:A.B(x)) ==> (lam x:A. f`x) = f
+
+\idx{lam_theI} (!!x. x:A ==> EX! y. Q(x,y)) ==> EX h. ALL x:A. Q(x, h`x)
+
+\idx{restrict_conv} a : A ==> restrict(f,A) ` a = f`a
+\idx{restrict_type} [| !!x. x:A ==> f`x: B(x) |] ==>
+ restrict(f,A) : (PROD x:A.B(x))
+
+\idx{fun_empty} 0: 0->0
+\idx{fun_single} \{<a,b>\} : \{a\} -> \{b\}
+
+\idx{fun_disjoint_Un} [| f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g) : (A Un C) -> (B Un D)
+
+\idx{fun_disjoint_apply1} [| a:A; f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g)`a = f`a
+
+\idx{fun_disjoint_apply2} [| c:C; f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g)`c = g`c
+
+
+%%% SIMPDATA.ML
+
+ a\in a & \bimp & False\\
+ a\in \emptyset & \bimp & False\\
+ a \in A \union B & \bimp & a\in A \disj a\in B\\
+ a \in A \inter B & \bimp & a\in A \conj a\in B\\
+ a \in A-B & \bimp & a\in A \conj \neg (a\in B)\\
+ a \in {\tt cons}(b,B) & \bimp & a=b \disj a\in B\\
+ i \in {\tt succ}(j) & \bimp & i=j \disj i\in j\\
+ A\in \bigcup(C) & \bimp & (\exists B. B\in C \conj A\in B)\\
+ A\in \bigcap(C) & \bimp & (\forall B. B\in C \imp A\in B)
+ \qquad (\exists B. B\in C)\\
+ a \in {\tt Collect}(A,P) & \bimp & a\in A \conj P(a)\\
+ b \in {\tt RepFun}(A,f) & \bimp & (\exists x. x\in A \conj b=f(x))
+
+equalities.ML perm.ML plus.ML nat.ML
+----------------------------------------------------------------
+equalities.ML
+
+\idx{Int_absorb} A Int A = A
+\idx{Int_commute} A Int B = B Int A
+\idx{Int_assoc} (A Int B) Int C = A Int (B Int C)
+\idx{Int_Un_distrib} (A Un B) Int C = (A Int C) Un (B Int C)
+
+\idx{Un_absorb} A Un A = A
+\idx{Un_commute} A Un B = B Un A
+\idx{Un_assoc} (A Un B) Un C = A Un (B Un C)
+\idx{Un_Int_distrib} (A Int B) Un C = (A Un C) Int (B Un C)
+
+\idx{Diff_cancel} A-A = 0
+\idx{Diff_disjoint} A Int (B-A) = 0
+\idx{Diff_partition} A<=B ==> A Un (B-A) = B
+\idx{double_complement} [| A<=B; B<= C |] ==> (B - (C-A)) = A
+\idx{Diff_Un} A - (B Un C) = (A-B) Int (A-C)
+\idx{Diff_Int} A - (B Int C) = (A-B) Un (A-C)
+
+\idx{Union_Un_distrib} Union(A Un B) = Union(A) Un Union(B)
+\idx{Inter_Un_distrib} [| a:A; b:B |] ==>
+ Inter(A Un B) = Inter(A) Int Inter(B)
+
+\idx{Int_Union_RepFun} A Int Union(B) = (UN C:B. A Int C)
+
+\idx{Un_Inter_RepFun} b:B ==>
+ A Un Inter(B) = (INT C:B. A Un C)
+
+\idx{SUM_Un_distrib1} (SUM x:A Un B. C(x)) =
+ (SUM x:A. C(x)) Un (SUM x:B. C(x))
+
+\idx{SUM_Un_distrib2} (SUM x:C. A(x) Un B(x)) =
+ (SUM x:C. A(x)) Un (SUM x:C. B(x))
+
+\idx{SUM_Int_distrib1} (SUM x:A Int B. C(x)) =
+ (SUM x:A. C(x)) Int (SUM x:B. C(x))
+
+\idx{SUM_Int_distrib2} (SUM x:C. A(x) Int B(x)) =
+ (SUM x:C. A(x)) Int (SUM x:C. B(x))
+
+
+----------------------------------------------------------------
+perm.ML
+
+\idx{comp_def}
+ r O s == \{xz : domain(s)*range(r) .
+ EX x y z. xz=<x,z> & <x,y>:s & <y,z>:r\}),
+\idx{id_def} (*the identity function for A*)
+ id(A) == (lam x:A. x)),
+\idx{inj_def}
+ inj(A,B) ==
+ \{ f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x\}),
+\idx{surj_def}
+ surj(A,B) == \{ f: A->B . ALL y:B. EX x:A. f`x=y\}),
+\idx{bij_def}
+ bij(A,B) == inj(A,B) Int surj(A,B))
+
+
+\idx{surj_is_fun} f: surj(A,B) ==> f: A->B
+\idx{fun_is_surj} f : Pi(A,B) ==> f: surj(A,range(f))
+
+\idx{inj_is_fun} f: inj(A,B) ==> f: A->B
+\idx{inj_equality} [| <a,b>:f; <c,b>:f; f: inj(A,B) |] ==> a=c
+
+\idx{bij_is_fun} f: bij(A,B) ==> f: A->B
+
+\idx{inj_converse_surj} f: inj(A,B) ==> converse(f): surj(range(f), A)
+
+\idx{left_inverse} [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a
+\idx{right_inverse} [| f: inj(A,B); b: range(f) |] ==>
+ f`(converse(f)`b) = b
+
+\idx{inj_converse_inj} f: inj(A,B) ==> converse(f): inj(range(f), A)
+\idx{bij_converse_bij} f: bij(A,B) ==> converse(f): bij(B,A)
+
+\idx{comp_type} [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C
+\idx{comp_assoc} (r O s) O t = r O (s O t)
+
+\idx{left_comp_id} r<=A*B ==> id(B) O r = r
+\idx{right_comp_id} r<=A*B ==> r O id(A) = r
+
+\idx{comp_func} [| g: A->B; f: B->C |] ==> (f O g) : A->C
+\idx{comp_func_apply} [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)
+
+\idx{comp_inj} [| g: inj(A,B); f: inj(B,C) |] ==> (f O g) : inj(A,C)
+\idx{comp_surj} [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)
+\idx{comp_bij} [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)
+
+\idx{left_comp_inverse} f: inj(A,B) ==> converse(f) O f = id(A)
+\idx{right_comp_inverse} f: surj(A,B) ==> f O converse(f) = id(B)
+
+\idx{bij_disjoint_Un}
+ [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==>
+ (f Un g) : bij(A Un C, B Un D)
+
+\idx{restrict_bij} [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)
+
+
+----------------------------------------------------------------
+plus.ML
+
+\idx{plus_def} A+B == \{0\}*A Un \{\{0\}\}*B
+\idx{Inl_def} Inl(a) == < 0 ,a>
+\idx{Inr_def} Inr(b) == <\{0\},b>
+\idx{when_def} when(u,c,d) ==
+ THE y. EX z.(u=Inl(z) & y=c(z)) | (u=Inr(z) & y=d(z))
+
+\idx{plus_InlI} a : A ==> Inl(a) : A+B
+\idx{plus_InrI} b : B ==> Inr(b) : A+B
+
+\idx{Inl_inject} Inl(a) = Inl(b) ==> a=b
+\idx{Inr_inject} Inr(a) = Inr(b) ==> a=b
+\idx{Inl_neq_Inr} Inl(a)=Inr(b) ==> P
+
+\idx{plusE2} u: A+B ==> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))
+
+\idx{when_Inl_conv} when(Inl(a),c,d) = c(a)
+\idx{when_Inr_conv} when(Inr(b),c,d) = d(b)
+
+\idx{when_type} [| u: A+B;
+ !!x. x: A ==> c(x): C(Inl(x));
+ !!y. y: B ==> d(y): C(Inr(y))
+ |] ==> when(u,c,d) : C(u)
+
+
+----------------------------------------------------------------
+nat.ML
+
+
+\idx{nat_def} nat == lfp(lam r: Pow(Inf). \{0\} Un RepFun(r,succ))
+\idx{nat_case_def} nat_case(n,a,b) ==
+ THE y. n=0 & y=a | (EX x. n=succ(x) & y=b(x))
+\idx{nat_rec_def} nat_rec(k,a,b) ==
+ transrec(nat, k, %n f. nat_case(n, a, %m. b(m, f`m)))
+
+\idx{nat_0_I} 0 : nat
+\idx{nat_succ_I} n : nat ==> succ(n) : nat
+
+\idx{nat_induct}
+ [| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x))
+ |] ==> P(n)
+
+\idx{nat_case_0_conv} nat_case(0,a,b) = a
+\idx{nat_case_succ_conv} nat_case(succ(m),a,b) = b(m)
+
+\idx{nat_case_type}
+ [| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m))
+ |] ==> nat_case(n,a,b) : C(n)
+
+\idx{nat_rec_0_conv} nat_rec(0,a,b) = a
+\idx{nat_rec_succ_conv} m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))
+
+\idx{nat_rec_type}
+ [| n: nat;
+ a: C(0);
+ !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m))
+ |] ==> nat_rec(n,a,b) : C(n)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/ZF.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,1571 @@
+%% $Id$
+%%%See grant/bra/Lib/ZF.tex for lfp figure
+\chapter{Zermelo-Fraenkel set theory}
+The directory~\ttindexbold{ZF} implements Zermelo-Fraenkel set
+theory~\cite{halmos60,suppes72} as an extension of~\ttindex{FOL}, classical
+first-order logic. The theory includes a collection of derived natural
+deduction rules, for use with Isabelle's classical reasoning module. Much
+of it is based on the work of No\"el~\cite{noel}. The theory has the {\ML}
+identifier \ttindexbold{ZF.thy}. However, many further theories
+are defined, introducing the natural numbers, etc.
+
+A tremendous amount of set theory has been formally developed, including
+the basic properties of relations, functions and ordinals. Significant
+results have been proved, such as the Schr\"oder-Bernstein Theorem and the
+Recursion Theorem. General methods have been developed for solving
+recursion equations over monotonic functors; these have been applied to
+yield constructions of lists and trees. Thus, we may even regard set
+theory as a computational logic. It admits recursive definitions of
+functions and types. It has similarities with Martin-L\"of type theory,
+although of course it is classical.
+
+Because {\ZF} is an extension of {\FOL}, it provides the same packages,
+namely \ttindex{hyp_subst_tac}, the simplifier, and the classical reasoning
+module. The main simplification set is called \ttindexbold{ZF_ss}.
+Several classical rule sets are defined, including \ttindexbold{lemmas_cs},
+\ttindexbold{upair_cs} and~\ttindexbold{ZF_cs}. See the files on directory
+{\tt ZF} for details.
+
+
+\section{Which version of axiomatic set theory?}
+Resolution theorem provers can work in set theory, using the
+Bernays-G\"odel axiom system~(BG) because it is
+finite~\cite{boyer86,quaife92}. {\ZF} does not have a finite axiom system
+(because of its Axiom Scheme of Replacement) and is therefore unsuitable
+for classical resolution. Since Isabelle has no difficulty with axiom
+schemes, we may adopt either axiom system.
+
+These two theories differ in their treatment of {\bf classes}, which are
+collections that are ``too big'' to be sets. The class of all sets,~$V$,
+cannot be a set without admitting Russell's Paradox. In BG, both classes
+and sets are individuals; $x\in V$ expresses that $x$ is a set. In {\ZF}, all
+variables denote sets; classes are identified with unary predicates. The
+two systems define essentially the same sets and classes, with similar
+properties. In particular, a class cannot belong to another class (let
+alone a set).
+
+Modern set theorists tend to prefer {\ZF} because they are mainly concerned
+with sets, rather than classes. BG requires tiresome proofs that various
+collections are sets; for instance, showing $x\in\{x\}$ requires showing that
+$x$ is a set. {\ZF} does not have this problem.
+
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{0} & $i$ & empty set\\
+ \idx{cons} & $[i,i]\To i$ & finite set constructor\\
+ \idx{Upair} & $[i,i]\To i$ & unordered pairing\\
+ \idx{Pair} & $[i,i]\To i$ & ordered pairing\\
+ \idx{Inf} & $i$ & infinite set\\
+ \idx{Pow} & $i\To i$ & powerset\\
+ \idx{Union} \idx{Inter} & $i\To i$ & set union/intersection \\
+ \idx{split} & $[i, [i,i]\To i] \To i$ & generalized projection\\
+ \idx{fst} \idx{snd} & $i\To i$ & projections\\
+ \idx{converse}& $i\To i$ & converse of a relation\\
+ \idx{succ} & $i\To i$ & successor\\
+ \idx{Collect} & $[i,i\To o]\To i$ & separation\\
+ \idx{Replace} & $[i, [i,i]\To o] \To i$ & replacement\\
+ \idx{PrimReplace} & $[i, [i,i]\To o] \To i$ & primitive replacement\\
+ \idx{RepFun} & $[i, i\To i] \To i$ & functional replacement\\
+ \idx{Pi} \idx{Sigma} & $[i,i\To i]\To i$ & general product/sum\\
+ \idx{domain} & $i\To i$ & domain of a relation\\
+ \idx{range} & $i\To i$ & range of a relation\\
+ \idx{field} & $i\To i$ & field of a relation\\
+ \idx{Lambda} & $[i, i\To i]\To i$ & $\lambda$-abstraction\\
+ \idx{restrict}& $[i, i] \To i$ & restriction of a function\\
+ \idx{The} & $[i\To o]\To i$ & definite description\\
+ \idx{if} & $[o,i,i]\To i$ & conditional\\
+ \idx{Ball} \idx{Bex} & $[i, i\To o]\To o$ & bounded quantifiers
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{center}
+\indexbold{*"`"`}
+\indexbold{*"-"`"`}
+\indexbold{*"`}
+\indexbold{*"-}
+\indexbold{*":}
+\indexbold{*"<"=}
+\begin{tabular}{rrrr}
+ \it symbol & \it meta-type & \it precedence & \it description \\
+ \tt `` & $[i,i]\To i$ & Left 90 & image \\
+ \tt -`` & $[i,i]\To i$ & Left 90 & inverse image \\
+ \tt ` & $[i,i]\To i$ & Left 90 & application \\
+ \idx{Int} & $[i,i]\To i$ & Left 70 & intersection ($\inter$) \\
+ \idx{Un} & $[i,i]\To i$ & Left 65 & union ($\union$) \\
+ \tt - & $[i,i]\To i$ & Left 65 & set difference ($-$) \\[1ex]
+ \tt: & $[i,i]\To o$ & Left 50 & membership ($\in$) \\
+ \tt <= & $[i,i]\To o$ & Left 50 & subset ($\subseteq$)
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+\caption{Constants of {\ZF}} \label{ZF-constants}
+\end{figure}
+
+
+\section{The syntax of set theory}
+The language of set theory, as studied by logicians, has no constants. The
+traditional axioms merely assert the existence of empty sets, unions,
+powersets, etc.; this would be intolerable for practical reasoning. The
+Isabelle theory declares constants for primitive sets. It also extends
+{\tt FOL} with additional syntax for finite sets, ordered pairs,
+comprehension, general union/intersection, general sums/products, and
+bounded quantifiers. In most other respects, Isabelle implements precisely
+Zermelo-Fraenkel set theory.
+
+Figure~\ref{ZF-constants} lists the constants and infixes of~\ZF, while
+Figure~\ref{ZF-trans} presents the syntax translations. Finally,
+Figure~\ref{ZF-syntax} presents the full grammar for set theory, including
+the constructs of \FOL.
+
+Set theory does not use polymorphism. All terms in {\ZF} have type~{\it
+i}, which is the type of individuals and lies in class {\it logic}.
+The type of first-order formulae,
+remember, is~{\it o}.
+
+Infix operators include union and intersection ($A\union B$ and $A\inter
+B$), and the subset and membership relations. Note that $a$\verb|~:|$b$ is
+translated to \verb|~(|$a$:$b$\verb|)|. The union and intersection
+operators ($\bigcup A$ and $\bigcap A$) form the union or intersection of a
+set of sets; $\bigcup A$ means the same as $\bigcup@{x\in A}x$. Of these
+operators, only $\bigcup A$ is primitive.
+
+The constant \ttindexbold{Upair} constructs unordered pairs; thus {\tt
+Upair($A$,$B$)} denotes the set~$\{A,B\}$ and {\tt Upair($A$,$A$)} denotes
+the singleton~$\{A\}$. As usual in {\ZF}, general union is used to define
+binary union. The Isabelle version goes on to define the constant
+\ttindexbold{cons}:
+\begin{eqnarray*}
+ A\cup B & \equiv & \bigcup({\tt Upair}(A,B)) \\
+ {\tt cons}(a,B) & \equiv & {\tt Upair}(a,a) \union B
+\end{eqnarray*}
+The {\tt\{\ldots\}} notation abbreviates finite sets constructed in the
+obvious manner using~{\tt cons} and~$\emptyset$ (the empty set):
+\begin{eqnarray*}
+ \{a,b,c\} & \equiv & {\tt cons}(a,{\tt cons}(b,{\tt cons}(c,\emptyset)))
+\end{eqnarray*}
+
+The constant \ttindexbold{Pair} constructs ordered pairs, as in {\tt
+Pair($a$,$b$)}. Ordered pairs may also be written within angle brackets,
+as {\tt<$a$,$b$>}.
+
+In {\ZF}, a function is a set of pairs. A {\ZF} function~$f$ is simply an
+individual as far as Isabelle is concerned: its Isabelle type is~$i$, not
+say $i\To i$. The infix operator~{\tt`} denotes the application of a
+function set to its argument; we must write~$f{\tt`}x$, not~$f(x)$. The
+syntax for image is~$f{\tt``}A$ and that for inverse image is~$f{\tt-``}A$.
+
+
+\begin{figure}
+\indexbold{*"-">}
+\indexbold{*"*}
+\begin{center} \tt\frenchspacing
+\begin{tabular}{rrr}
+ \it external & \it internal & \it description \\
+ $a$ \ttilde: $b$ & \ttilde($a$ : $b$) & \rm negated membership\\
+ \{$a@1$, $\ldots$, $a@n$\} & cons($a@1$,$\cdots$,cons($a@n$,0)) &
+ \rm finite set \\
+ <$a$, $b$> & Pair($a$,$b$) & \rm ordered pair \\
+ <$a$, $b$, $c$> & <$a$, <$b$, $c$>> & \rm nested pairs (any depth) \\
+ \{$x$:$A . P[x]$\} & Collect($A$,$\lambda x.P[x]$) &
+ \rm separation \\
+ \{$y . x$:$A$, $Q[x,y]$\} & Replace($A$,$\lambda x\,y.Q[x,y]$) &
+ \rm replacement \\
+ \{$b[x] . x$:$A$\} & RepFun($A$,$\lambda x.b[x]$) &
+ \rm functional replacement \\
+ \idx{INT} $x$:$A . B[x]$ & Inter(\{$B[x] . x$:$A$\}) &
+ \rm general intersection \\
+ \idx{UN} $x$:$A . B[x]$ & Union(\{$B[x] . x$:$A$\}) &
+ \rm general union \\
+ \idx{PROD} $x$:$A . B[x]$ & Pi($A$,$\lambda x.B[x]$) &
+ \rm general product \\
+ \idx{SUM} $x$:$A . B[x]$ & Sigma($A$,$\lambda x.B[x]$) &
+ \rm general sum \\
+ $A$ -> $B$ & Pi($A$,$\lambda x.B$) &
+ \rm function space \\
+ $A$ * $B$ & Sigma($A$,$\lambda x.B$) &
+ \rm binary product \\
+ \idx{THE} $x . P[x]$ & The($\lambda x.P[x]$) &
+ \rm definite description \\
+ \idx{lam} $x$:$A . b[x]$ & Lambda($A$,$\lambda x.b[x]$) &
+ \rm $\lambda$-abstraction\\[1ex]
+ \idx{ALL} $x$:$A . P[x]$ & Ball($A$,$\lambda x.P[x]$) &
+ \rm bounded $\forall$ \\
+ \idx{EX} $x$:$A . P[x]$ & Bex($A$,$\lambda x.P[x]$) &
+ \rm bounded $\exists$
+\end{tabular}
+\end{center}
+\caption{Translations for {\ZF}} \label{ZF-trans}
+\end{figure}
+
+
+\begin{figure}
+\dquotes
+\[\begin{array}{rcl}
+ term & = & \hbox{expression of type~$i$} \\
+ & | & "\{ " term\; ("," term)^* " \}" \\
+ & | & "< " term ", " term " >" \\
+ & | & "\{ " id ":" term " . " formula " \}" \\
+ & | & "\{ " id " . " id ":" term "," formula " \}" \\
+ & | & "\{ " term " . " id ":" term " \}" \\
+ & | & term " `` " term \\
+ & | & term " -`` " term \\
+ & | & term " ` " term \\
+ & | & term " * " term \\
+ & | & term " Int " term \\
+ & | & term " Un " term \\
+ & | & term " - " term \\
+ & | & term " -> " term \\
+ & | & "THE~~" id " . " formula\\
+ & | & "lam~~" id ":" term " . " term \\
+ & | & "INT~~" id ":" term " . " term \\
+ & | & "UN~~~" id ":" term " . " term \\
+ & | & "PROD~" id ":" term " . " term \\
+ & | & "SUM~~" id ":" term " . " term \\[2ex]
+ formula & = & \hbox{expression of type~$o$} \\
+ & | & term " : " term \\
+ & | & term " <= " term \\
+ & | & term " = " term \\
+ & | & "\ttilde\ " formula \\
+ & | & formula " \& " formula \\
+ & | & formula " | " formula \\
+ & | & formula " --> " formula \\
+ & | & formula " <-> " formula \\
+ & | & "ALL " id ":" term " . " formula \\
+ & | & "EX~~" id ":" term " . " formula \\
+ & | & "ALL~" id~id^* " . " formula \\
+ & | & "EX~~" id~id^* " . " formula \\
+ & | & "EX!~" id~id^* " . " formula
+ \end{array}
+\]
+\caption{Full grammar for {\ZF}} \label{ZF-syntax}
+\end{figure}
+
+
+\section{Binding operators}
+The constant \ttindexbold{Collect} constructs sets by the principle of {\bf
+ separation}. The syntax for separation is \hbox{\tt\{$x$:$A$.$P[x]$\}},
+where $P[x]$ is a formula that may contain free occurrences of~$x$. It
+abbreviates the set {\tt Collect($A$,$\lambda x.P$[x])}, which consists of
+all $x\in A$ that satisfy~$P[x]$. Note that {\tt Collect} is an
+unfortunate choice of name: some set theories adopt a set-formation
+principle, related to replacement, called collection.
+
+The constant \ttindexbold{Replace} constructs sets by the principle of {\bf
+ replacement}. The syntax for replacement is
+\hbox{\tt\{$y$.$x$:$A$,$Q[x,y]$\}}. It denotes the set {\tt
+ Replace($A$,$\lambda x\,y.Q$[x,y])} consisting of all $y$ such that there
+exists $x\in A$ satisfying~$Q[x,y]$. The Replacement Axiom has the
+condition that $Q$ must be single-valued over~$A$: for all~$x\in A$ there
+exists at most one $y$ satisfying~$Q[x,y]$. A single-valued binary
+predicate is also called a {\bf class function}.
+
+The constant \ttindexbold{RepFun} expresses a special case of replacement,
+where $Q[x,y]$ has the form $y=b[x]$. Such a $Q$ is trivially
+single-valued, since it is just the graph of the meta-level
+function~$\lambda x.b[x]$. The syntax is \hbox{\tt\{$b[x]$.$x$:$A$\}},
+denoting set {\tt RepFun($A$,$\lambda x.b[x]$)} of all $b[x]$ for~$x\in A$.
+This is analogous to the \ML{} functional {\tt map}, since it applies a
+function to every element of a set.
+
+\indexbold{*INT}\indexbold{*UN}
+General unions and intersections of families, namely $\bigcup@{x\in A}B[x]$ and
+$\bigcap@{x\in A}B[x]$, are written \hbox{\tt UN $x$:$A$.$B[x]$} and
+\hbox{\tt INT $x$:$A$.$B[x]$}. Their meaning is expressed using {\tt
+RepFun} as
+\[ \bigcup(\{B[x]. x\in A\}) \qquad\hbox{and}\qquad
+ \bigcap(\{B[x]. x\in A\}).
+\]
+General sums $\sum@{x\in A}B[x]$ and products $\prod@{x\in A}B[x]$ can be
+constructed in set theory, where $B[x]$ is a family of sets over~$A$. They
+have as special cases $A\times B$ and $A\to B$, where $B$ is simply a set.
+This is similar to the situation in Constructive Type Theory (set theory
+has ``dependent sets'') and calls for similar syntactic conventions. The
+constants~\ttindexbold{Sigma} and~\ttindexbold{Pi} construct general sums and
+products. Instead of {\tt Sigma($A$,$B$)} and {\tt Pi($A$,$B$)} we may write
+\hbox{\tt SUM $x$:$A$.$B[x]$} and \hbox{\tt PROD $x$:$A$.$B[x]$}.
+\indexbold{*SUM}\indexbold{*PROD}%
+The special cases as \hbox{\tt$A$*$B$} and \hbox{\tt$A$->$B$} abbreviate
+general sums and products over a constant family.\footnote{Unlike normal
+infix operators, {\tt*} and {\tt->} merely define abbreviations; there are
+no constants~{\tt op~*} and~\hbox{\tt op~->}.} Isabelle accepts these
+abbreviations in parsing and uses them whenever possible for printing.
+
+\indexbold{*THE}
+As mentioned above, whenever the axioms assert the existence and uniqueness
+of a set, Isabelle's set theory declares a constant for that set. These
+constants can express the {\bf definite description} operator~$\iota
+x.P[x]$, which stands for the unique~$a$ satisfying~$P[a]$, if such exists.
+Since all terms in {\ZF} denote something, a description is always
+meaningful, but we do not know its value unless $P[x]$ defines it uniquely.
+Using the constant~\ttindexbold{The}, we may write descriptions as {\tt
+ The($\lambda x.P[x]$)} or use the syntax \hbox{\tt THE $x$.$P[x]$}.
+
+\indexbold{*lam}
+Function sets may be written in $\lambda$-notation; $\lambda x\in A.b[x]$
+stands for the set of all pairs $\pair{x,b[x]}$ for $x\in A$. In order for
+this to be a set, the function's domain~$A$ must be given. Using the
+constant~\ttindexbold{Lambda}, we may express function sets as {\tt
+Lambda($A$,$\lambda x.b[x]$)} or use the syntax \hbox{\tt lam $x$:$A$.$b[x]$}.
+
+Isabelle's set theory defines two {\bf bounded quantifiers}:
+\begin{eqnarray*}
+ \forall x\in A.P[x] &\hbox{which abbreviates}& \forall x. x\in A\imp P[x] \\
+ \exists x\in A.P[x] &\hbox{which abbreviates}& \exists x. x\in A\conj P[x]
+\end{eqnarray*}
+The constants~\ttindexbold{Ball} and~\ttindexbold{Bex} are defined
+accordingly. Instead of {\tt Ball($A$,$P$)} and {\tt Bex($A$,$P$)} we may
+write
+\hbox{\tt ALL $x$:$A$.$P[x]$} and \hbox{\tt EX $x$:$A$.$P[x]$}.
+
+
+%%%% zf.thy
+
+\begin{figure}
+\begin{ttbox}
+\idx{Ball_def} Ball(A,P) == ALL x. x:A --> P(x)
+\idx{Bex_def} Bex(A,P) == EX x. x:A & P(x)
+
+\idx{subset_def} A <= B == ALL x:A. x:B
+\idx{extension} A = B <-> A <= B & B <= A
+
+\idx{union_iff} A : Union(C) <-> (EX B:C. A:B)
+\idx{power_set} A : Pow(B) <-> A <= B
+\idx{foundation} A=0 | (EX x:A. ALL y:x. ~ y:A)
+
+\idx{replacement} (ALL x:A. ALL y z. P(x,y) & P(x,z) --> y=z) ==>
+ b : PrimReplace(A,P) <-> (EX x:A. P(x,b))
+\subcaption{The Zermelo-Fraenkel Axioms}
+
+\idx{Replace_def} Replace(A,P) ==
+ PrimReplace(A, %x y. (EX!z.P(x,z)) & P(x,y))
+\idx{RepFun_def} RepFun(A,f) == \{y . x:A, y=f(x)\}
+\idx{the_def} The(P) == Union(\{y . x:\{0\}, P(y)\})
+\idx{if_def} if(P,a,b) == THE z. P & z=a | ~P & z=b
+\idx{Collect_def} Collect(A,P) == \{y . x:A, x=y & P(x)\}
+\idx{Upair_def} Upair(a,b) ==
+ \{y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)\}
+\subcaption{Consequences of replacement}
+
+\idx{Inter_def} Inter(A) == \{ x:Union(A) . ALL y:A. x:y\}
+\idx{Un_def} A Un B == Union(Upair(A,B))
+\idx{Int_def} A Int B == Inter(Upair(A,B))
+\idx{Diff_def} A - B == \{ x:A . ~(x:B) \}
+\subcaption{Union, intersection, difference}
+
+\idx{cons_def} cons(a,A) == Upair(a,a) Un A
+\idx{succ_def} succ(i) == cons(i,i)
+\idx{infinity} 0:Inf & (ALL y:Inf. succ(y): Inf)
+\subcaption{Finite and infinite sets}
+\end{ttbox}
+\caption{Rules and axioms of {\ZF}} \label{ZF-rules}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{Pair_def} <a,b> == \{\{a,a\}, \{a,b\}\}
+\idx{split_def} split(p,c) == THE y. EX a b. p=<a,b> & y=c(a,b)
+\idx{fst_def} fst(A) == split(p, %x y.x)
+\idx{snd_def} snd(A) == split(p, %x y.y)
+\idx{Sigma_def} Sigma(A,B) == UN x:A. UN y:B(x). \{<x,y>\}
+\subcaption{Ordered pairs and Cartesian products}
+
+\idx{converse_def} converse(r) == \{z. w:r, EX x y. w=<x,y> & z=<y,x>\}
+\idx{domain_def} domain(r) == \{x. w:r, EX y. w=<x,y>\}
+\idx{range_def} range(r) == domain(converse(r))
+\idx{field_def} field(r) == domain(r) Un range(r)
+\idx{image_def} r `` A == \{y : range(r) . EX x:A. <x,y> : r\}
+\idx{vimage_def} r -`` A == converse(r)``A
+\subcaption{Operations on relations}
+
+\idx{lam_def} Lambda(A,b) == \{<x,b(x)> . x:A\}
+\idx{apply_def} f`a == THE y. <a,y> : f
+\idx{Pi_def} Pi(A,B) == \{f: Pow(Sigma(A,B)). ALL x:A. EX! y. <x,y>: f\}
+\idx{restrict_def} restrict(f,A) == lam x:A.f`x
+\subcaption{Functions and general product}
+\end{ttbox}
+\caption{Further definitions of {\ZF}} \label{ZF-defs}
+\end{figure}
+
+
+%%%% zf.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{ballI} [| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)
+\idx{bspec} [| ALL x:A. P(x); x: A |] ==> P(x)
+\idx{ballE} [| ALL x:A. P(x); P(x) ==> Q; ~ x:A ==> Q |] ==> Q
+
+\idx{ball_cong} [| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==>
+ (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))
+
+\idx{bexI} [| P(x); x: A |] ==> EX x:A. P(x)
+\idx{bexCI} [| ALL x:A. ~P(x) ==> P(a); a: A |] ==> EX x:A.P(x)
+\idx{bexE} [| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q
+
+\idx{bex_cong} [| A=A'; !!x. x:A' ==> P(x) <-> P'(x) |] ==>
+ (EX x:A. P(x)) <-> (EX x:A'. P'(x))
+\subcaption{Bounded quantifiers}
+
+\idx{subsetI} (!!x.x:A ==> x:B) ==> A <= B
+\idx{subsetD} [| A <= B; c:A |] ==> c:B
+\idx{subsetCE} [| A <= B; ~(c:A) ==> P; c:B ==> P |] ==> P
+\idx{subset_refl} A <= A
+\idx{subset_trans} [| A<=B; B<=C |] ==> A<=C
+
+\idx{equalityI} [| A <= B; B <= A |] ==> A = B
+\idx{equalityD1} A = B ==> A<=B
+\idx{equalityD2} A = B ==> B<=A
+\idx{equalityE} [| A = B; [| A<=B; B<=A |] ==> P |] ==> P
+\subcaption{Subsets and extensionality}
+
+\idx{emptyE} a:0 ==> P
+\idx{empty_subsetI} 0 <= A
+\idx{equals0I} [| !!y. y:A ==> False |] ==> A=0
+\idx{equals0D} [| A=0; a:A |] ==> P
+
+\idx{PowI} A <= B ==> A : Pow(B)
+\idx{PowD} A : Pow(B) ==> A<=B
+\subcaption{The empty set; power sets}
+\end{ttbox}
+\caption{Basic derived rules for {\ZF}} \label{ZF-lemmas1}
+\end{figure}
+
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{ReplaceI} [| x: A; P(x,b); !!y. P(x,y) ==> y=b |] ==>
+ b : \{y. x:A, P(x,y)\}
+
+\idx{ReplaceE} [| b : \{y. x:A, P(x,y)\};
+ !!x. [| x: A; P(x,b); ALL y. P(x,y)-->y=b |] ==> R
+ |] ==> R
+
+\idx{RepFunI} [| a : A |] ==> f(a) : \{f(x). x:A\}
+\idx{RepFunE} [| b : \{f(x). x:A\};
+ !!x.[| x:A; b=f(x) |] ==> P |] ==> P
+
+\idx{separation} a : \{x:A. P(x)\} <-> a:A & P(a)
+\idx{CollectI} [| a:A; P(a) |] ==> a : \{x:A. P(x)\}
+\idx{CollectE} [| a : \{x:A. P(x)\}; [| a:A; P(a) |] ==> R |] ==> R
+\idx{CollectD1} a : \{x:A. P(x)\} ==> a:A
+\idx{CollectD2} a : \{x:A. P(x)\} ==> P(a)
+\end{ttbox}
+\caption{Replacement and separation} \label{ZF-lemmas2}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{UnionI} [| B: C; A: B |] ==> A: Union(C)
+\idx{UnionE} [| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R
+
+\idx{InterI} [| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)
+\idx{InterD} [| A : Inter(C); B : C |] ==> A : B
+\idx{InterE} [| A : Inter(C); A:B ==> R; ~ B:C ==> R |] ==> R
+
+\idx{UN_I} [| a: A; b: B(a) |] ==> b: (UN x:A. B(x))
+\idx{UN_E} [| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R
+ |] ==> R
+
+\idx{INT_I} [| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))
+\idx{INT_E} [| b : (INT x:A. B(x)); a: A |] ==> b : B(a)
+\end{ttbox}
+\caption{General Union and Intersection} \label{ZF-lemmas3}
+\end{figure}
+
+
+\section{The Zermelo-Fraenkel axioms}
+The axioms appear in Figure~\ref{ZF-rules}. They resemble those
+presented by Suppes~\cite{suppes72}. Most of the theory consists of
+definitions. In particular, bounded quantifiers and the subset relation
+appear in other axioms. Object-level quantifiers and implications have
+been replaced by meta-level ones wherever possible, to simplify use of the
+axioms. See the file \ttindexbold{ZF/zf.thy} for details.
+
+The traditional replacement axiom asserts
+\[ y \in {\tt PrimReplace}(A,P) \bimp (\exists x\in A. P(x,y)) \]
+subject to the condition that $P(x,y)$ is single-valued for all~$x\in A$.
+The Isabelle theory defines \ttindex{Replace} to apply
+\ttindexbold{PrimReplace} to the single-valued part of~$P$, namely
+\[ (\exists!z.P(x,z)) \conj P(x,y). \]
+Thus $y\in {\tt Replace}(A,P)$ if and only if there is some~$x$ such that
+$P(x,-)$ holds uniquely for~$y$. Because the equivalence is unconditional,
+{\tt Replace} is much easier to use than {\tt PrimReplace}; it defines the
+same set, if $P(x,y)$ is single-valued. The nice syntax for replacement
+expands to {\tt Replace}.
+
+Other consequences of replacement include functional replacement
+(\ttindexbold{RepFun}) and definite descriptions (\ttindexbold{The}).
+Axioms for separation (\ttindexbold{Collect}) and unordered pairs
+(\ttindexbold{Upair}) are traditionally assumed, but they actually follow
+from replacement~\cite[pages 237--8]{suppes72}.
+
+The definitions of general intersection, etc., are straightforward. Note
+the definition of \ttindex{cons}, which underlies the finite set notation.
+The axiom of infinity gives us a set that contains~0 and is closed under
+successor (\ttindexbold{succ}). Although this set is not uniquely defined,
+the theory names it (\ttindexbold{Inf}) in order to simplify the
+construction of the natural numbers.
+
+Further definitions appear in Figure~\ref{ZF-defs}. Ordered pairs are
+defined in the standard way, $\pair{a,b}\equiv\{\{a\},\{a,b\}\}$. Recall
+that \ttindexbold{Sigma}$(A,B)$ generalizes the Cartesian product of two
+sets. It is defined to be the union of all singleton sets
+$\{\pair{x,y}\}$, for $x\in A$ and $y\in B(x)$. This is a typical usage of
+general union.
+
+The projections involve definite descriptions. The \ttindex{split}
+operation is like the similar operation in Martin-L\"of Type Theory, and is
+often easier to use than \ttindex{fst} and~\ttindex{snd}. It is defined
+using a description for convenience, but could equivalently be defined by
+\begin{ttbox}
+split(p,c) == c(fst(p),snd(p))
+\end{ttbox}
+Operations on relations include converse, domain, range, and image. The
+set ${\tt Pi}(A,B)$ generalizes the space of functions between two sets.
+Note the simple definitions of $\lambda$-abstraction (using
+\ttindex{RepFun}) and application (using a definite description). The
+function \ttindex{restrict}$(f,A)$ has the same values as~$f$, but only
+over the domain~$A$.
+
+No axiom of choice is provided. It is traditional to include this axiom
+only where it is needed --- mainly in the theory of cardinal numbers, which
+Isabelle does not formalize at present.
+
+
+\section{From basic lemmas to function spaces}
+Faced with so many definitions, it is essential to prove lemmas. Even
+trivial theorems like $A\inter B=B\inter A$ would be difficult to prove
+from the definitions alone. Isabelle's set theory derives many rules using
+a natural deduction style. Ideally, a natural deduction rule should
+introduce or eliminate just one operator, but this is not always practical.
+For most operators, we may forget its definition and use its derived rules
+instead.
+
+\subsection{Fundamental lemmas}
+Figure~\ref{ZF-lemmas1} presents the derived rules for the most basic
+operators. The rules for the bounded quantifiers resemble those for the
+ordinary quantifiers, but note that \ttindex{BallE} uses a negated
+assumption in the style of Isabelle's classical module. The congruence rules
+\ttindex{ball_cong} and \ttindex{bex_cong} are required by Isabelle's
+simplifier, but have few other uses. Congruence rules must be specially
+derived for all binding operators, and henceforth will not be shown.
+
+Figure~\ref{ZF-lemmas1} also shows rules for the subset and equality
+relations (proof by extensionality), and rules about the empty set and the
+power set operator.
+
+Figure~\ref{ZF-lemmas2} presents rules for replacement and separation.
+The rules for \ttindex{Replace} and \ttindex{RepFun} are much simpler than
+comparable rules for {\tt PrimReplace} would be. The principle of
+separation is proved explicitly, although most proofs should use the
+natural deduction rules for \ttindex{Collect}. The elimination rule
+\ttindex{CollectE} is equivalent to the two destruction rules
+\ttindex{CollectD1} and \ttindex{CollectD2}, but each rule is suited to
+particular circumstances. Although too many rules can be confusing, there
+is no reason to aim for a minimal set of rules. See the file
+\ttindexbold{ZF/zf.ML} for a complete listing.
+
+Figure~\ref{ZF-lemmas3} presents rules for general union and intersection.
+The empty intersection should be undefined. We cannot have
+$\bigcap(\emptyset)=V$ because $V$, the universal class, is not a set. All
+expressions denote something in {\ZF} set theory; the definition of
+intersection implies $\bigcap(\emptyset)=\emptyset$, but this value is
+arbitrary. The rule \ttindexbold{InterI} must have a premise to exclude
+the empty intersection. Some of the laws governing intersections require
+similar premises.
+
+
+%%% upair.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{pairing} a:Upair(b,c) <-> (a=b | a=c)
+\idx{UpairI1} a : Upair(a,b)
+\idx{UpairI2} b : Upair(a,b)
+\idx{UpairE} [| a : Upair(b,c); a = b ==> P; a = c ==> P |] ==> P
+\subcaption{Unordered pairs}
+
+\idx{UnI1} c : A ==> c : A Un B
+\idx{UnI2} c : B ==> c : A Un B
+\idx{UnCI} (~c : B ==> c : A) ==> c : A Un B
+\idx{UnE} [| c : A Un B; c:A ==> P; c:B ==> P |] ==> P
+
+\idx{IntI} [| c : A; c : B |] ==> c : A Int B
+\idx{IntD1} c : A Int B ==> c : A
+\idx{IntD2} c : A Int B ==> c : B
+\idx{IntE} [| c : A Int B; [| c:A; c:B |] ==> P |] ==> P
+
+\idx{DiffI} [| c : A; ~ c : B |] ==> c : A - B
+\idx{DiffD1} c : A - B ==> c : A
+\idx{DiffD2} [| c : A - B; c : B |] ==> P
+\idx{DiffE} [| c : A - B; [| c:A; ~ c:B |] ==> P |] ==> P
+\subcaption{Union, intersection, difference}
+\end{ttbox}
+\caption{Unordered pairs and their consequences} \label{ZF-upair1}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{consI1} a : cons(a,B)
+\idx{consI2} a : B ==> a : cons(b,B)
+\idx{consCI} (~ a:B ==> a=b) ==> a: cons(b,B)
+\idx{consE} [| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P
+
+\idx{singletonI} a : \{a\}
+\idx{singletonE} [| a : \{b\}; a=b ==> P |] ==> P
+\subcaption{Finite and singleton sets}
+
+\idx{succI1} i : succ(i)
+\idx{succI2} i : j ==> i : succ(j)
+\idx{succCI} (~ i:j ==> i=j) ==> i: succ(j)
+\idx{succE} [| i : succ(j); i=j ==> P; i:j ==> P |] ==> P
+\idx{succ_neq_0} [| succ(n)=0 |] ==> P
+\idx{succ_inject} succ(m) = succ(n) ==> m=n
+\subcaption{The successor function}
+
+\idx{the_equality} [| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a
+\idx{theI} EX! x. P(x) ==> P(THE x. P(x))
+
+\idx{if_P} P ==> if(P,a,b) = a
+\idx{if_not_P} ~P ==> if(P,a,b) = b
+
+\idx{mem_anti_sym} [| a:b; b:a |] ==> P
+\idx{mem_anti_refl} a:a ==> P
+\subcaption{Descriptions; non-circularity}
+\end{ttbox}
+\caption{Finite sets and their consequences} \label{ZF-upair2}
+\end{figure}
+
+
+\subsection{Unordered pairs and finite sets}
+Figure~\ref{ZF-upair1} presents the principle of unordered pairing, along
+with its derived rules. Binary union and intersection are defined in terms
+of ordered pairs, and set difference is included for completeness. The
+rule \ttindexbold{UnCI} is useful for classical reasoning about unions,
+like {\tt disjCI}\@; it supersedes \ttindexbold{UnI1} and
+\ttindexbold{UnI2}, but these rules are often easier to work with. For
+intersection and difference we have both elimination and destruction rules.
+Again, there is no reason to provide a minimal rule set.
+
+Figure~\ref{ZF-upair2} is concerned with finite sets. It presents rules
+for~\ttindex{cons}, the finite set constructor, and rules for singleton
+sets. Because the successor function is defined in terms of~{\tt cons},
+its derived rules appear here.
+
+Definite descriptions (\ttindex{THE}) are defined in terms of the singleton
+set $\{0\}$, but their derived rules fortunately hide this. The
+rule~\ttindex{theI} can be difficult to apply, because $\Var{P}$ must be
+instantiated correctly. However, \ttindex{the_equality} does not have this
+problem and the files contain many examples of its use.
+
+Finally, the impossibility of having both $a\in b$ and $b\in a$
+(\ttindex{mem_anti_sym}) is proved by applying the axiom of foundation to
+the set $\{a,b\}$. The impossibility of $a\in a$ is a trivial consequence.
+
+See the file \ttindexbold{ZF/upair.ML} for full details.
+
+
+%%% subset.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{Union_upper} B:A ==> B <= Union(A)
+\idx{Union_least} [| !!x. x:A ==> x<=C |] ==> Union(A) <= C
+
+\idx{Inter_lower} B:A ==> Inter(A) <= B
+\idx{Inter_greatest} [| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)
+
+\idx{Un_upper1} A <= A Un B
+\idx{Un_upper2} B <= A Un B
+\idx{Un_least} [| A<=C; B<=C |] ==> A Un B <= C
+
+\idx{Int_lower1} A Int B <= A
+\idx{Int_lower2} A Int B <= B
+\idx{Int_greatest} [| C<=A; C<=B |] ==> C <= A Int B
+
+\idx{Diff_subset} A-B <= A
+\idx{Diff_contains} [| C<=A; C Int B = 0 |] ==> C <= A-B
+
+\idx{Collect_subset} Collect(A,P) <= A
+\end{ttbox}
+\caption{Subset and lattice properties} \label{ZF-subset}
+\end{figure}
+
+
+\subsection{Subset and lattice properties}
+Figure~\ref{ZF-subset} shows that the subset relation is a complete
+lattice. Unions form least upper bounds; non-empty intersections form
+greatest lower bounds. A few other laws involving subsets are included.
+See the file \ttindexbold{ZF/subset.ML}.
+
+%%% pair.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{Pair_inject1} <a,b> = <c,d> ==> a=c
+\idx{Pair_inject2} <a,b> = <c,d> ==> b=d
+\idx{Pair_inject} [| <a,b> = <c,d>; [| a=c; b=d |] ==> P |] ==> P
+\idx{Pair_neq_0} <a,b>=0 ==> P
+
+\idx{fst} fst(<a,b>) = a
+\idx{snd} snd(<a,b>) = b
+\idx{split} split(<a,b>, %x y.c(x,y)) = c(a,b)
+
+\idx{SigmaI} [| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)
+
+\idx{SigmaE} [| c: Sigma(A,B);
+ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P |] ==> P
+
+\idx{SigmaE2} [| <a,b> : Sigma(A,B);
+ [| a:A; b:B(a) |] ==> P |] ==> P
+\end{ttbox}
+\caption{Ordered pairs; projections; general sums} \label{ZF-pair}
+\end{figure}
+
+
+\subsection{Ordered pairs}
+Figure~\ref{ZF-pair} presents the rules governing ordered pairs,
+projections and general sums. File \ttindexbold{ZF/pair.ML} contains the
+full (and tedious) proof that $\{\{a\},\{a,b\}\}$ functions as an ordered
+pair. This property is expressed as two destruction rules,
+\ttindexbold{Pair_inject1} and \ttindexbold{Pair_inject2}, and equivalently
+as the elimination rule \ttindexbold{Pair_inject}.
+
+Note the rule \ttindexbold{Pair_neq_0}, which asserts
+$\pair{a,b}\neq\emptyset$. This is no arbitrary property of
+$\{\{a\},\{a,b\}\}$, but one that we can reasonably expect to hold for any
+encoding of ordered pairs. It turns out to be useful for constructing
+Lisp-style S-expressions in set theory.
+
+The natural deduction rules \ttindexbold{SigmaI} and \ttindexbold{SigmaE}
+assert that \ttindex{Sigma}$(A,B)$ consists of all pairs of the form
+$\pair{x,y}$, for $x\in A$ and $y\in B(x)$. The rule \ttindexbold{SigmaE2}
+merely states that $\pair{a,b}\in {\tt Sigma}(A,B)$ implies $a\in A$ and
+$b\in B(a)$.
+
+
+%%% domrange.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{domainI} <a,b>: r ==> a : domain(r)
+\idx{domainE} [| a : domain(r); !!y. <a,y>: r ==> P |] ==> P
+\idx{domain_subset} domain(Sigma(A,B)) <= A
+
+\idx{rangeI} <a,b>: r ==> b : range(r)
+\idx{rangeE} [| b : range(r); !!x. <x,b>: r ==> P |] ==> P
+\idx{range_subset} range(A*B) <= B
+
+\idx{fieldI1} <a,b>: r ==> a : field(r)
+\idx{fieldI2} <a,b>: r ==> b : field(r)
+\idx{fieldCI} (~ <c,a>:r ==> <a,b>: r) ==> a : field(r)
+
+\idx{fieldE} [| a : field(r);
+ !!x. <a,x>: r ==> P;
+ !!x. <x,a>: r ==> P
+ |] ==> P
+
+\idx{field_subset} field(A*A) <= A
+\subcaption{Domain, range and field of a Relation}
+
+\idx{imageI} [| <a,b>: r; a:A |] ==> b : r``A
+\idx{imageE} [| b: r``A; !!x.[| <x,b>: r; x:A |] ==> P |] ==> P
+
+\idx{vimageI} [| <a,b>: r; b:B |] ==> a : r-``B
+\idx{vimageE} [| a: r-``B; !!x.[| <a,x>: r; x:B |] ==> P |] ==> P
+\subcaption{Image and inverse image}
+\end{ttbox}
+\caption{Relations} \label{ZF-domrange}
+\end{figure}
+
+
+\subsection{Relations}
+Figure~\ref{ZF-domrange} presents rules involving relations, which are sets
+of ordered pairs. The converse of a relation~$r$ is the set of all pairs
+$\pair{y,x}$ such that $\pair{x,y}\in r$; if $r$ is a function, then
+{\ttindex{converse}$(r)$} is its inverse. The rules for the domain
+operation, \ttindex{domainI} and~\ttindex{domainE}, assert that
+\ttindex{domain}$(r)$ consists of every element~$a$ such that $r$ contains
+some pair of the form~$\pair{x,y}$. The range operation is similar, and
+the field of a relation is merely the union of its domain and range. Note
+that image and inverse image are generalizations of range and domain,
+respectively. See the file
+\ttindexbold{ZF/domrange.ML} for derivations of the rules.
+
+
+%%% func.ML
+
+\begin{figure}
+\begin{ttbox}
+\idx{fun_is_rel} f: Pi(A,B) ==> f <= Sigma(A,B)
+
+\idx{apply_equality} [| <a,b>: f; f: Pi(A,B) |] ==> f`a = b
+\idx{apply_equality2} [| <a,b>: f; <a,c>: f; f: Pi(A,B) |] ==> b=c
+
+\idx{apply_type} [| f: Pi(A,B); a:A |] ==> f`a : B(a)
+\idx{apply_Pair} [| f: Pi(A,B); a:A |] ==> <a,f`a>: f
+\idx{apply_iff} f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b
+
+\idx{fun_extension} [| f : Pi(A,B); g: Pi(A,D);
+ !!x. x:A ==> f`x = g`x |] ==> f=g
+
+\idx{domain_type} [| <a,b> : f; f: Pi(A,B) |] ==> a : A
+\idx{range_type} [| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)
+
+\idx{Pi_type} [| f: A->C; !!x. x:A ==> f`x: B(x) |] ==> f: Pi(A,B)
+\idx{domain_of_fun} f: Pi(A,B) ==> domain(f)=A
+\idx{range_of_fun} f: Pi(A,B) ==> f: A->range(f)
+
+\idx{restrict} a : A ==> restrict(f,A) ` a = f`a
+\idx{restrict_type} [| !!x. x:A ==> f`x: B(x) |] ==>
+ restrict(f,A) : Pi(A,B)
+
+\idx{lamI} a:A ==> <a,b(a)> : (lam x:A. b(x))
+\idx{lamE} [| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P
+ |] ==> P
+
+\idx{lam_type} [| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A.b(x)) : Pi(A,B)
+
+\idx{beta} a : A ==> (lam x:A.b(x)) ` a = b(a)
+\idx{eta} f : Pi(A,B) ==> (lam x:A. f`x) = f
+
+\idx{lam_theI} (!!x. x:A ==> EX! y. Q(x,y)) ==> EX h. ALL x:A. Q(x, h`x)
+\end{ttbox}
+\caption{Functions and $\lambda$-abstraction} \label{ZF-func1}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{fun_empty} 0: 0->0
+\idx{fun_single} \{<a,b>\} : \{a\} -> \{b\}
+
+\idx{fun_disjoint_Un} [| f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g) : (A Un C) -> (B Un D)
+
+\idx{fun_disjoint_apply1} [| a:A; f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g)`a = f`a
+
+\idx{fun_disjoint_apply2} [| c:C; f: A->B; g: C->D; A Int C = 0 |] ==>
+ (f Un g)`c = g`c
+\end{ttbox}
+\caption{Constructing functions from smaller sets} \label{ZF-func2}
+\end{figure}
+
+
+\subsection{Functions}
+Functions, represented by graphs, are notoriously difficult to reason
+about. The file \ttindexbold{ZF/func.ML} derives many rules, which overlap
+more than they ought. One day these rules will be tidied up; this section
+presents only the more important ones.
+
+Figure~\ref{ZF-func1} presents the basic properties of \ttindex{Pi}$(A,B)$,
+the generalized function space. For example, if $f$ is a function and
+$\pair{a,b}\in f$, then $f`a=b$ (\ttindex{apply_equality}). Two functions
+are equal provided they have equal domains and deliver equals results
+(\ttindex{fun_extension}).
+
+By \ttindex{Pi_type}, a function typing of the form $f\in A\to C$ can be
+refined to the dependent typing $f\in\prod@{x\in A}B(x)$, given a suitable
+family of sets $\{B(x)\}@{x\in A}$. Conversely, by \ttindex{range_of_fun},
+any dependent typing can be flattened to yield a function type of the form
+$A\to C$; here, $C={\tt range}(f)$.
+
+Among the laws for $\lambda$-abstraction, \ttindex{lamI} and \ttindex{lamE}
+describe the graph of the generated function, while \ttindex{beta} and
+\ttindex{eta} are the standard conversions. We essentially have a
+dependently-typed $\lambda$-calculus.
+
+Figure~\ref{ZF-func2} presents some rules that can be used to construct
+functions explicitly. We start with functions consisting of at most one
+pair, and may form the union of two functions provided their domains are
+disjoint.
+
+
+\begin{figure}
+\begin{center}
+\begin{tabular}{rrr}
+ \it name &\it meta-type & \it description \\
+ \idx{id} & $i$ & identity function \\
+ \idx{inj} & $[i,i]\To i$ & injective function space\\
+ \idx{surj} & $[i,i]\To i$ & surjective function space\\
+ \idx{bij} & $[i,i]\To i$ & bijective function space
+ \\[1ex]
+ \idx{1} & $i$ & $\{\emptyset\}$ \\
+ \idx{bool} & $i$ & the set $\{\emptyset,1\}$ \\
+ \idx{cond} & $[i,i,i]\To i$ & conditional for {\tt bool}
+ \\[1ex]
+ \idx{Inl}~~\idx{Inr} & $i\To i$ & injections\\
+ \idx{case} & $[i,i\To i,i\To i]\To i$ & conditional for $+$
+ \\[1ex]
+ \idx{nat} & $i$ & set of natural numbers \\
+ \idx{nat_case}& $[i,i,i\To i]\To i$ & conditional for $nat$\\
+ \idx{rec} & $[i,i,[i,i]\To i]\To i$ & recursor for $nat$
+ \\[1ex]
+ \idx{list} & $i\To i$ & lists over some set\\
+ \idx{list_case} & $[i, i, [i,i]\To i] \To i$ & conditional for $list(A)$ \\
+ \idx{list_rec} & $[i, i, [i,i,i]\To i] \To i$ & recursor for $list(A)$ \\
+ \idx{map} & $[i\To i, i] \To i$ & mapping functional\\
+ \idx{length} & $i\To i$ & length of a list\\
+ \idx{rev} & $i\To i$ & reverse of a list\\
+ \idx{flat} & $i\To i$ & flatting a list of lists\\
+\end{tabular}
+\end{center}
+\subcaption{Constants}
+
+\begin{center}
+\indexbold{*"+}
+\index{#*@{\tt\#*}|bold}
+\index{*div|bold}
+\index{*mod|bold}
+\index{#+@{\tt\#+}|bold}
+\index{#-@{\tt\#-}|bold}
+\begin{tabular}{rrrr}
+ \idx{O} & $[i,i]\To i$ & Right 60 & composition ($\circ$) \\
+ \tt + & $[i,i]\To i$ & Right 65 & disjoint union \\
+ \tt \#* & $[i,i]\To i$ & Left 70 & multiplication \\
+ \tt div & $[i,i]\To i$ & Left 70 & division\\
+ \tt mod & $[i,i]\To i$ & Left 70 & modulus\\
+ \tt \#+ & $[i,i]\To i$ & Left 65 & addition\\
+ \tt \#- & $[i,i]\To i$ & Left 65 & subtraction\\
+ \tt \@ & $[i,i]\To i$ & Right 60 & append for lists
+\end{tabular}
+\end{center}
+\subcaption{Infixes}
+\caption{Further constants for {\ZF}} \label{ZF-further-constants}
+\end{figure}
+
+
+\begin{figure}
+\begin{ttbox}
+\idx{Int_absorb} A Int A = A
+\idx{Int_commute} A Int B = B Int A
+\idx{Int_assoc} (A Int B) Int C = A Int (B Int C)
+\idx{Int_Un_distrib} (A Un B) Int C = (A Int C) Un (B Int C)
+
+\idx{Un_absorb} A Un A = A
+\idx{Un_commute} A Un B = B Un A
+\idx{Un_assoc} (A Un B) Un C = A Un (B Un C)
+\idx{Un_Int_distrib} (A Int B) Un C = (A Un C) Int (B Un C)
+
+\idx{Diff_cancel} A-A = 0
+\idx{Diff_disjoint} A Int (B-A) = 0
+\idx{Diff_partition} A<=B ==> A Un (B-A) = B
+\idx{double_complement} [| A<=B; B<= C |] ==> (B - (C-A)) = A
+\idx{Diff_Un} A - (B Un C) = (A-B) Int (A-C)
+\idx{Diff_Int} A - (B Int C) = (A-B) Un (A-C)
+
+\idx{Union_Un_distrib} Union(A Un B) = Union(A) Un Union(B)
+\idx{Inter_Un_distrib} [| a:A; b:B |] ==>
+ Inter(A Un B) = Inter(A) Int Inter(B)
+
+\idx{Int_Union_RepFun} A Int Union(B) = (UN C:B. A Int C)
+
+\idx{Un_Inter_RepFun} b:B ==>
+ A Un Inter(B) = (INT C:B. A Un C)
+
+\idx{SUM_Un_distrib1} (SUM x:A Un B. C(x)) =
+ (SUM x:A. C(x)) Un (SUM x:B. C(x))
+
+\idx{SUM_Un_distrib2} (SUM x:C. A(x) Un B(x)) =
+ (SUM x:C. A(x)) Un (SUM x:C. B(x))
+
+\idx{SUM_Int_distrib1} (SUM x:A Int B. C(x)) =
+ (SUM x:A. C(x)) Int (SUM x:B. C(x))
+
+\idx{SUM_Int_distrib2} (SUM x:C. A(x) Int B(x)) =
+ (SUM x:C. A(x)) Int (SUM x:C. B(x))
+\end{ttbox}
+\caption{Equalities} \label{zf-equalities}
+\end{figure}
+
+\begin{figure}
+\begin{ttbox}
+\idx{comp_def} r O s == \{xz : domain(s)*range(r) .
+ EX x y z. xz=<x,z> & <x,y>:s & <y,z>:r\}
+\idx{id_def} id(A) == (lam x:A. x)
+\idx{inj_def} inj(A,B) == \{ f: A->B. ALL w:A. ALL x:A. f`w=f`x --> w=x\}
+\idx{surj_def} surj(A,B) == \{ f: A->B . ALL y:B. EX x:A. f`x=y\}
+\idx{bij_def} bij(A,B) == inj(A,B) Int surj(A,B)
+\subcaption{Definitions}
+
+\idx{left_inverse} [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a
+\idx{right_inverse} [| f: inj(A,B); b: range(f) |] ==>
+ f`(converse(f)`b) = b
+
+\idx{inj_converse_inj} f: inj(A,B) ==> converse(f): inj(range(f), A)
+\idx{bij_converse_bij} f: bij(A,B) ==> converse(f): bij(B,A)
+
+\idx{comp_type} [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C
+\idx{comp_assoc} (r O s) O t = r O (s O t)
+
+\idx{left_comp_id} r<=A*B ==> id(B) O r = r
+\idx{right_comp_id} r<=A*B ==> r O id(A) = r
+
+\idx{comp_func} [| g:A->B; f:B->C |] ==> (f O g):A->C
+\idx{comp_func_apply} [| g:A->B; f:B->C; a:A |] ==> (f O g)`a = f`(g`a)
+
+\idx{comp_inj} [| g:inj(A,B); f:inj(B,C) |] ==> (f O g):inj(A,C)
+\idx{comp_surj} [| g:surj(A,B); f:surj(B,C) |] ==> (f O g):surj(A,C)
+\idx{comp_bij} [| g:bij(A,B); f:bij(B,C) |] ==> (f O g):bij(A,C)
+
+\idx{left_comp_inverse} f: inj(A,B) ==> converse(f) O f = id(A)
+\idx{right_comp_inverse} f: surj(A,B) ==> f O converse(f) = id(B)
+
+\idx{bij_disjoint_Un}
+ [| f: bij(A,B); g: bij(C,D); A Int C = 0; B Int D = 0 |] ==>
+ (f Un g) : bij(A Un C, B Un D)
+
+\idx{restrict_bij} [| f:inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)
+\end{ttbox}
+\caption{Permutations} \label{zf-perm}
+\end{figure}
+
+\begin{figure}
+\begin{ttbox}
+\idx{one_def} 1 == succ(0)
+\idx{bool_def} bool == {0,1}
+\idx{cond_def} cond(b,c,d) == if(b=1,c,d)
+
+\idx{sum_def} A+B == \{0\}*A Un \{1\}*B
+\idx{Inl_def} Inl(a) == <0,a>
+\idx{Inr_def} Inr(b) == <1,b>
+\idx{case_def} case(u,c,d) == split(u, %y z. cond(y, d(z), c(z)))
+\subcaption{Definitions}
+
+\idx{bool_1I} 1 : bool
+\idx{bool_0I} 0 : bool
+
+\idx{boolE} [| c: bool; P(1); P(0) |] ==> P(c)
+\idx{cond_1} cond(1,c,d) = c
+\idx{cond_0} cond(0,c,d) = d
+
+\idx{sum_InlI} a : A ==> Inl(a) : A+B
+\idx{sum_InrI} b : B ==> Inr(b) : A+B
+
+\idx{Inl_inject} Inl(a)=Inl(b) ==> a=b
+\idx{Inr_inject} Inr(a)=Inr(b) ==> a=b
+\idx{Inl_neq_Inr} Inl(a)=Inr(b) ==> P
+
+\idx{sumE2} u: A+B ==> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))
+
+\idx{case_Inl} case(Inl(a),c,d) = c(a)
+\idx{case_Inr} case(Inr(b),c,d) = d(b)
+\end{ttbox}
+\caption{Booleans and disjoint unions} \label{zf-sum}
+\end{figure}
+
+\begin{figure}
+\begin{ttbox}
+\idx{nat_def} nat == lfp(lam r: Pow(Inf). \{0\} Un \{succ(x). x:r\}
+
+\idx{nat_case_def} nat_case(n,a,b) ==
+ THE y. n=0 & y=a | (EX x. n=succ(x) & y=b(x))
+
+\idx{rec_def} rec(k,a,b) ==
+ transrec(k, %n f. nat_case(n, a, %m. b(m, f`m)))
+
+\idx{add_def} m#+n == rec(m, n, %u v.succ(v))
+\idx{diff_def} m#-n == rec(n, m, %u v. rec(v, 0, %x y.x))
+\idx{mult_def} m#*n == rec(m, 0, %u v. n #+ v)
+\idx{mod_def} m mod n == transrec(m, %j f. if(j:n, j, f`(j#-n)))
+\idx{quo_def} m div n == transrec(m, %j f. if(j:n, 0, succ(f`(j#-n))))
+\subcaption{Definitions}
+
+\idx{nat_0I} 0 : nat
+\idx{nat_succI} n : nat ==> succ(n) : nat
+
+\idx{nat_induct}
+ [| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x))
+ |] ==> P(n)
+
+\idx{nat_case_0} nat_case(0,a,b) = a
+\idx{nat_case_succ} nat_case(succ(m),a,b) = b(m)
+
+\idx{rec_0} rec(0,a,b) = a
+\idx{rec_succ} rec(succ(m),a,b) = b(m, rec(m,a,b))
+
+\idx{mult_type} [| m:nat; n:nat |] ==> m #* n : nat
+\idx{mult_0} 0 #* n = 0
+\idx{mult_succ} succ(m) #* n = n #+ (m #* n)
+\idx{mult_commute} [| m:nat; n:nat |] ==> m #* n = n #* m
+\idx{add_mult_dist}
+ [| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)
+\idx{mult_assoc}
+ [| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)
+
+\idx{mod_quo_equality}
+ [| 0:n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m
+\end{ttbox}
+\caption{The natural numbers} \label{zf-nat}
+\end{figure}
+
+\begin{figure}\underscoreon %%because @ is used here
+\begin{ttbox}
+\idx{list_def} list(A) == lfp(univ(A), %X. {0} Un A*X)
+
+\idx{list_case_def} list_case(l,c,h) ==
+ THE z. l=0 & z=c | (EX x y. l = <x,y> & z=h(x,y))
+
+\idx{list_rec_def} list_rec(l,c,h) ==
+ Vrec(l, %l g.list_case(l, c, %x xs. h(x, xs, g`xs)))
+
+\idx{map_def} map(f,l) == list_rec(l, 0, %x xs r. <f(x), r>)
+\idx{length_def} length(l) == list_rec(l, 0, %x xs r. succ(r))
+\idx{app_def} xs@ys == list_rec(xs, ys, %x xs r. <x,r>)
+\idx{rev_def} rev(l) == list_rec(l, 0, %x xs r. r @ <x,0>)
+\idx{flat_def} flat(ls) == list_rec(ls, 0, %l ls r. l @ r)
+\subcaption{Definitions}
+
+\idx{list_0I} 0 : list(A)
+\idx{list_PairI} [| a: A; l: list(A) |] ==> <a,l> : list(A)
+
+\idx{list_induct}
+ [| l: list(A);
+ P(0);
+ !!x y. [| x: A; y: list(A); P(y) |] ==> P(<x,y>)
+ |] ==> P(l)
+
+\idx{list_case_0} list_case(0,c,h) = c
+\idx{list_case_Pair} list_case(<a,l>, c, h) = h(a,l)
+
+\idx{list_rec_0} list_rec(0,c,h) = c
+\idx{list_rec_Pair} list_rec(<a,l>, c, h) = h(a, l, list_rec(l,c,h))
+
+\idx{map_ident} l: list(A) ==> map(%u.u, l) = l
+\idx{map_compose} l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)
+\idx{map_app_distrib} xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)
+\idx{map_type}
+ [| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)
+\idx{map_flat}
+ ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))
+\end{ttbox}
+\caption{Lists} \label{zf-list}
+\end{figure}
+
+\section{Further developments}
+The next group of developments is complex and extensive, and only
+highlights can be covered here. Figure~\ref{ZF-further-constants} lists
+some of the further constants and infixes that are declared in the various
+theory extensions.
+
+Figure~\ref{zf-sum} defines $\{0,1\}$ as a set of booleans, with a
+conditional operator. It also defines the disjoint union of two sets, with
+injections and a case analysis operator. See files
+\ttindexbold{ZF/bool.ML} and \ttindexbold{ZF/sum.ML}.
+
+Monotonicity properties of most of the set-forming operations are proved.
+These are useful for applying the Knaster-Tarski Fixedpoint Theorem.
+See file \ttindexbold{ZF/mono.ML}.
+
+Figure~\ref{zf-equalities} presents commutative, associative, distributive,
+and idempotency laws of union and intersection, along with other equations.
+See file \ttindexbold{ZF/equalities.ML}.
+
+Figure~\ref{zf-perm} defines composition and injective, surjective and
+bijective function spaces, with proofs of many of their properties.
+See file \ttindexbold{ZF/perm.ML}.
+
+Figure~\ref{zf-nat} presents the natural numbers, with induction and a
+primitive recursion operator. See file \ttindexbold{ZF/nat.ML}. File
+\ttindexbold{ZF/arith.ML} develops arithmetic on the natural numbers. It
+defines addition, multiplication, subtraction, division, and remainder,
+proving the theorem $a \bmod b + (a/b)\times b = a$. Division and
+remainder are defined by repeated subtraction, which requires well-founded
+rather than primitive recursion.
+
+Figure~\ref{zf-list} presents defines the set of lists over~$A$, ${\tt
+list}(A)$ as the least solution of the equation ${\tt list}(A)\equiv \{0\}
+\union (A\times{\tt list}(A))$. Structural induction and recursion are
+derived, with some of the usual list functions. See file
+\ttindexbold{ZF/list.ML}.
+
+The constructions of the natural numbers and lists make use of a suite of
+operators for handling recursive definitions. The developments are
+summarized below:
+\begin{description}
+\item[\ttindexbold{ZF/lfp.ML}]
+proves the Knaster-Tarski Fixedpoint Theorem in the lattice of subsets of a
+set. The file defines a least fixedpoint operator with corresponding
+induction rules. These are used repeatedly in the sequel to define sets
+inductively. As a simple application, the file contains a short proof of
+the Schr\"oder-Bernstein Theorem.
+
+\item[\ttindexbold{ZF/trancl.ML}]
+defines the transitive closure of a relation (as a least fixedpoint).
+
+\item[\ttindexbold{ZF/wf.ML}]
+proves the Well-Founded Recursion Theorem, using an elegant
+approach of Tobias Nipkow. This theorem permits general recursive
+definitions within set theory.
+
+\item[\ttindexbold{ZF/ordinal.ML}]
+defines the notions of transitive set and ordinal number. It derives
+transfinite induction.
+
+\item[\ttindexbold{ZF/epsilon.ML}]
+derives $\epsilon$-induction and $\epsilon$-recursion, which are
+generalizations of transfinite induction. It also defines
+\ttindexbold{rank}$(x)$, which is the least ordinal $\alpha$ such that $x$
+is constructed at stage $\alpha$ of the cumulative hierarchy (thus $x\in
+V@{\alpha+1}$).
+
+\item[\ttindexbold{ZF/univ.ML}]
+defines a ``universe'' ${\tt univ}(A)$, for constructing sets inductively.
+This set contains $A$ and the natural numbers. Vitally, it is
+closed under finite products:
+${\tt univ}(A)\times{\tt univ}(A)\subseteq{\tt univ}(A)$. This file also
+defines set theory's cumulative hierarchy, traditionally written $V@\alpha$
+where $\alpha$ is any ordinal.
+\end{description}
+
+
+\begin{figure}
+\begin{eqnarray*}
+ a\in a & \bimp & \bot\\
+ a\in \emptyset & \bimp & \bot\\
+ a \in A \union B & \bimp & a\in A \disj a\in B\\
+ a \in A \inter B & \bimp & a\in A \conj a\in B\\
+ a \in A-B & \bimp & a\in A \conj \neg (a\in B)\\
+ a \in {\tt cons}(b,B) & \bimp & a=b \disj a\in B\\
+ i \in {\tt succ}(j) & \bimp & i=j \disj i\in j\\
+ \pair{a,b}\in {\tt Sigma}(A,B)
+ & \bimp & a\in A \conj b\in B(a)\\
+ a \in {\tt Collect}(A,P) & \bimp & a\in A \conj P(a)\\
+ (\forall x \in A. \top) & \bimp & \top
+\end{eqnarray*}
+\caption{Rewrite rules for set theory} \label{ZF-simpdata}
+\end{figure}
+
+
+\section{Simplification rules}
+{\ZF} does not merely inherit simplification from \FOL, but instantiates
+the rewriting package new. File \ttindexbold{ZF/simpdata.ML} contains the
+details; here is a summary of the key differences:
+\begin{itemize}
+\item
+\ttindexbold{mk_rew_rules} is given as a function that can
+strip bounded universal quantifiers from a formula. For example, $\forall
+x\in A.f(x)=g(x)$ yields the conditional rewrite rule $x\in A \Imp
+f(x)=g(x)$.
+\item
+\ttindexbold{ZF_ss} contains congruence rules for all the operators of
+{\ZF}, including the binding operators. It contains all the conversion
+rules, such as \ttindex{fst} and \ttindex{snd}, as well as the
+rewrites shown in Figure~\ref{ZF-simpdata}.
+\item
+\ttindexbold{FOL_ss} is redeclared with the same {\FOL} rules as the
+previous version, so that it may still be used.
+\end{itemize}
+
+
+\section{The examples directory}
+This directory contains further developments in {\ZF} set theory. Here is
+an overview; see the files themselves for more details.
+\begin{description}
+\item[\ttindexbold{ZF/ex/misc.ML}]
+contains miscellaneous examples such as Cantor's Theorem and the
+``Composition of homomorphisms'' challenge.
+
+\item[\ttindexbold{ZF/ex/ramsey.ML}]
+proves the finite exponent 2 version of Ramsey's Theorem.
+
+\item[\ttindexbold{ZF/ex/bt.ML}]
+defines the recursive data structure ${\tt bt}(A)$, labelled binary trees.
+
+\item[\ttindexbold{ZF/ex/sexp.ML}]
+defines the set of Lisp $S$-expressions over~$A$, ${\tt sexp}(A)$. These
+are unlabelled binary trees whose leaves contain elements of $(A)$.
+
+\item[\ttindexbold{ZF/ex/term.ML}]
+defines a recursive data structure for terms and term lists.
+
+\item[\ttindexbold{ZF/ex/simult.ML}]
+defines primitives for solving mutually recursive equations over sets.
+It constructs sets of trees and forests as an example, including induction
+and recursion rules that handle the mutual recursion.
+
+\item[\ttindexbold{ZF/ex/finite.ML}]
+inductively defines a finite powerset operator.
+
+\item[\ttindexbold{ZF/ex/prop-log.ML}]
+proves soundness and completeness of propositional logic. This illustrates
+the main forms of induction.
+\end{description}
+
+
+\section{A proof about powersets}
+To demonstrate high-level reasoning about subsets, let us prove the equation
+${Pow(A)\cap Pow(B)}= Pow(A\cap B)$. Compared with first-order logic, set
+theory involves a maze of rules, and theorems have many different proofs.
+Attempting other proofs of the theorem might be instructive. This proof
+exploits the lattice properties of intersection. It also uses the
+monotonicity of the powerset operation, from {\tt ZF/mono.ML}:
+\begin{ttbox}
+\idx{Pow_mono} A<=B ==> Pow(A) <= Pow(B)
+\end{ttbox}
+We enter the goal and make the first step, which breaks the equation into
+two inclusions by extensionality:\index{equalityI}
+\begin{ttbox}
+goal ZF.thy "Pow(A Int B) = Pow(A) Int Pow(B)";
+{\out Level 0}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A Int B) = Pow(A) Int Pow(B)}
+by (resolve_tac [equalityI] 1);
+{\out Level 1}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A Int B) <= Pow(A) Int Pow(B)}
+{\out 2. Pow(A) Int Pow(B) <= Pow(A Int B)}
+\end{ttbox}
+Both inclusions could be tackled straightforwardly using {\tt subsetI}.
+A shorter proof results from noting that intersection forms the greatest
+lower bound:\index{*Int_greatest}
+\begin{ttbox}
+by (resolve_tac [Int_greatest] 1);
+{\out Level 2}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A Int B) <= Pow(A)}
+{\out 2. Pow(A Int B) <= Pow(B)}
+{\out 3. Pow(A) Int Pow(B) <= Pow(A Int B)}
+\end{ttbox}
+Subgoal~1 follows by applying the monotonicity of {\tt Pow} to $A\inter
+B\subseteq A$; subgoal~2 follows similarly:
+\index{*Int_lower1}\index{*Int_lower2}
+\begin{ttbox}
+by (resolve_tac [Int_lower1 RS Pow_mono] 1);
+{\out Level 3}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A Int B) <= Pow(B)}
+{\out 2. Pow(A) Int Pow(B) <= Pow(A Int B)}
+by (resolve_tac [Int_lower2 RS Pow_mono] 1);
+{\out Level 4}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A) Int Pow(B) <= Pow(A Int B)}
+\end{ttbox}
+We are left with the opposite inclusion, which we tackle in the
+straightforward way:\index{*subsetI}
+\begin{ttbox}
+by (resolve_tac [subsetI] 1);
+{\out Level 5}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. !!x. x : Pow(A) Int Pow(B) ==> x : Pow(A Int B)}
+\end{ttbox}
+The subgoal is to show $x\in {\tt Pow}(A\cap B)$ assuming $x\in{\tt
+Pow}(A)\cap {\tt Pow}(B)$. Eliminating this assumption produces two
+subgoals, since intersection is like conjunction.\index{*IntE}
+\begin{ttbox}
+by (eresolve_tac [IntE] 1);
+{\out Level 6}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. !!x. [| x : Pow(A); x : Pow(B) |] ==> x : Pow(A Int B)}
+\end{ttbox}
+The next step replaces the {\tt Pow} by the subset
+relation~($\subseteq$).\index{*PowI}
+\begin{ttbox}
+by (resolve_tac [PowI] 1);
+{\out Level 7}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. !!x. [| x : Pow(A); x : Pow(B) |] ==> x <= A Int B}
+\end{ttbox}
+We perform the same replacement in the assumptions:\index{*PowD}
+\begin{ttbox}
+by (REPEAT (dresolve_tac [PowD] 1));
+{\out Level 8}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. !!x. [| x <= A; x <= B |] ==> x <= A Int B}
+\end{ttbox}
+Here, $x$ is a lower bound of $A$ and~$B$, but $A\inter B$ is the greatest
+lower bound:\index{*Int_greatest}
+\begin{ttbox}
+by (resolve_tac [Int_greatest] 1);
+{\out Level 9}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. !!x. [| x <= A; x <= B |] ==> x <= A}
+{\out 2. !!x. [| x <= A; x <= B |] ==> x <= B}
+by (REPEAT (assume_tac 1));
+{\out Level 10}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out No subgoals!}
+\end{ttbox}
+We could have performed this proof in one step by applying
+\ttindex{fast_tac} with the classical rule set \ttindex{ZF_cs}. But we
+must add \ttindex{equalityI} as an introduction rule, since extensionality
+is not used by default:
+\begin{ttbox}
+choplev 0;
+{\out Level 0}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out 1. Pow(A Int B) = Pow(A) Int Pow(B)}
+by (fast_tac (ZF_cs addIs [equalityI]) 1);
+{\out Level 1}
+{\out Pow(A Int B) = Pow(A) Int Pow(B)}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\section{Monotonicity of the union operator}
+For another example, we prove that general union is monotonic:
+${C\subseteq D}$ implies $\bigcup(C)\subseteq \bigcup(D)$. To begin, we
+tackle the inclusion using \ttindex{subsetI}:
+\begin{ttbox}
+val [prem] = goal ZF.thy "C<=D ==> Union(C) <= Union(D)";
+{\out Level 0}
+{\out Union(C) <= Union(D)}
+{\out 1. Union(C) <= Union(D)}
+by (resolve_tac [subsetI] 1);
+{\out Level 1}
+{\out Union(C) <= Union(D)}
+{\out 1. !!x. x : Union(C) ==> x : Union(D)}
+\end{ttbox}
+Big union is like an existential quantifier --- the occurrence in the
+assumptions must be eliminated early, since it creates parameters.
+\index{*UnionE}
+\begin{ttbox}
+by (eresolve_tac [UnionE] 1);
+{\out Level 2}
+{\out Union(C) <= Union(D)}
+{\out 1. !!x B. [| x : B; B : C |] ==> x : Union(D)}
+\end{ttbox}
+Now we may apply \ttindex{UnionI}, which creates an unknown involving the
+parameters. To show $x\in \bigcup(D)$ it suffices to show that $x$ belongs
+to some element, say~$\Var{B2}(x,B)$, of~$D$.
+\begin{ttbox}
+by (resolve_tac [UnionI] 1);
+{\out Level 3}
+{\out Union(C) <= Union(D)}
+{\out 1. !!x B. [| x : B; B : C |] ==> ?B2(x,B) : D}
+{\out 2. !!x B. [| x : B; B : C |] ==> x : ?B2(x,B)}
+\end{ttbox}
+Combining \ttindex{subsetD} with the premise $C\subseteq D$ yields
+$\Var{a}\in C \Imp \Var{a}\in D$, which reduces subgoal~1:
+\begin{ttbox}
+by (resolve_tac [prem RS subsetD] 1);
+{\out Level 4}
+{\out Union(C) <= Union(D)}
+{\out 1. !!x B. [| x : B; B : C |] ==> ?B2(x,B) : C}
+{\out 2. !!x B. [| x : B; B : C |] ==> x : ?B2(x,B)}
+\end{ttbox}
+The rest is routine. Note how~$\Var{B2}(x,B)$ is instantiated.
+\begin{ttbox}
+by (assume_tac 1);
+{\out Level 5}
+{\out Union(C) <= Union(D)}
+{\out 1. !!x B. [| x : B; B : C |] ==> x : B}
+by (assume_tac 1);
+{\out Level 6}
+{\out Union(C) <= Union(D)}
+{\out No subgoals!}
+\end{ttbox}
+Again, \ttindex{fast_tac} with \ttindex{ZF_cs} can do this proof in one
+step, provided we somehow supply it with~{\tt prem}. We can either add
+this premise to the assumptions using \ttindex{cut_facts_tac}, or add
+\hbox{\tt prem RS subsetD} to \ttindex{ZF_cs} as an introduction rule.
+
+The file \ttindex{ZF/equalities.ML} has many similar proofs.
+Reasoning about general intersection can be difficult because of its anomalous
+behavior on the empty set. However, \ttindex{fast_tac} copes well with
+these. Here is a typical example:
+\begin{ttbox}
+a:C ==> (INT x:C. A(x) Int B(x)) = (INT x:C.A(x)) Int (INT x:C.B(x))
+\end{ttbox}
+In traditional notation this is
+\[ a\in C \,\Imp\, \bigcap@{x\in C} \Bigl(A(x) \inter B(x)\Bigr) =
+ \Bigl(\bigcap@{x\in C} A(x)\Bigr) \inter
+ \Bigl(\bigcap@{x\in C} B(x)\Bigr) \]
+
+\section{Low-level reasoning about functions}
+The derived rules {\tt lamI}, {\tt lamE}, {\tt lam_type}, {\tt beta}
+and {\tt eta} support reasoning about functions in a
+$\lambda$-calculus style. This is generally easier than regarding
+functions as sets of ordered pairs. But sometimes we must look at the
+underlying representation, as in the following proof
+of~\ttindex{fun_disjoint_apply1}. This states that if $f$ and~$g$ are
+functions with disjoint domains~$A$ and~$C$, and if $a\in A$, then
+$(f\union g)`a = f`a$:
+\begin{ttbox}
+val prems = goal ZF.thy
+ "[| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \ttback
+\ttback (f Un g)`a = f`a";
+{\out Level 0}
+{\out (f Un g) ` a = f ` a}
+{\out 1. (f Un g) ` a = f ` a}
+\end{ttbox}
+Using \ttindex{apply_equality}, we reduce the equality to reasoning about
+ordered pairs.
+\begin{ttbox}
+by (resolve_tac [apply_equality] 1);
+{\out Level 1}
+{\out (f Un g) ` a = f ` a}
+{\out 1. <a,f ` a> : f Un g}
+{\out 2. f Un g : (PROD x:?A. ?B(x))}
+\end{ttbox}
+We must show that the pair belongs to~$f$ or~$g$; by~\ttindex{UnI1} we
+choose~$f$:
+\begin{ttbox}
+by (resolve_tac [UnI1] 1);
+{\out Level 2}
+{\out (f Un g) ` a = f ` a}
+{\out 1. <a,f ` a> : f}
+{\out 2. f Un g : (PROD x:?A. ?B(x))}
+\end{ttbox}
+To show $\pair{a,f`a}\in f$ we use \ttindex{apply_Pair}, which is
+essentially the converse of \ttindex{apply_equality}:
+\begin{ttbox}
+by (resolve_tac [apply_Pair] 1);
+{\out Level 3}
+{\out (f Un g) ` a = f ` a}
+{\out 1. f : (PROD x:?A2. ?B2(x))}
+{\out 2. a : ?A2}
+{\out 3. f Un g : (PROD x:?A. ?B(x))}
+\end{ttbox}
+Using the premises $f\in A\to B$ and $a\in A$, we solve the two subgoals
+from \ttindex{apply_Pair}. Recall that a $\Pi$-set is merely a generalized
+function space, and observe that~{\tt?A2} is instantiated to~{\tt A}.
+\begin{ttbox}
+by (resolve_tac prems 1);
+{\out Level 4}
+{\out (f Un g) ` a = f ` a}
+{\out 1. a : A}
+{\out 2. f Un g : (PROD x:?A. ?B(x))}
+by (resolve_tac prems 1);
+{\out Level 5}
+{\out (f Un g) ` a = f ` a}
+{\out 1. f Un g : (PROD x:?A. ?B(x))}
+\end{ttbox}
+To construct functions of the form $f\union g$, we apply
+\ttindex{fun_disjoint_Un}:
+\begin{ttbox}
+by (resolve_tac [fun_disjoint_Un] 1);
+{\out Level 6}
+{\out (f Un g) ` a = f ` a}
+{\out 1. f : ?A3 -> ?B3}
+{\out 2. g : ?C3 -> ?D3}
+{\out 3. ?A3 Int ?C3 = 0}
+\end{ttbox}
+The remaining subgoals are instances of the premises. Again, observe how
+unknowns are instantiated:
+\begin{ttbox}
+by (resolve_tac prems 1);
+{\out Level 7}
+{\out (f Un g) ` a = f ` a}
+{\out 1. g : ?C3 -> ?D3}
+{\out 2. A Int ?C3 = 0}
+by (resolve_tac prems 1);
+{\out Level 8}
+{\out (f Un g) ` a = f ` a}
+{\out 1. A Int C = 0}
+by (resolve_tac prems 1);
+{\out Level 9}
+{\out (f Un g) ` a = f ` a}
+{\out No subgoals!}
+\end{ttbox}
+See the files \ttindex{ZF/func.ML} and \ttindex{ZF/wf.ML} for more
+examples of reasoning about functions.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/abstract.txt Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,8 @@
+Isabelle's Object-Logics. Report 286.
+
+This is the third of the Isabelle manuals. It covers Isabelle's built-in
+object logics: first-order logic (FOL), Zermelo-Fraenkel set theory (ZF),
+higher-order logic (HOL), the classical sequent calculus (LK), and
+constructive type theory (CTT). The final chapter discusses how to define
+new logics. This manual assumes familiarity with Report 280, Introduction
+to Isabelle.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/defining.tex Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,1557 @@
+%% $Id$
+\newcommand{\rmindex}[1]{{#1}\index{#1}\@}
+\newcommand{\mtt}[1]{\mbox{\tt #1}}
+\newcommand{\ttfct}[1]{\mathop{\mtt{#1}}\nolimits}
+\newcommand{\ttrel}[1]{\mathrel{\mtt{#1}}}
+\newcommand{\Constant}{\ttfct{Constant}}
+\newcommand{\Variable}{\ttfct{Variable}}
+\newcommand{\Appl}[1]{\ttfct{Appl}\mathopen{\mtt[}#1\mathclose{\mtt]}}
+
+
+
+\chapter{Defining Logics} \label{Defining-Logics}
+
+This chapter is intended for Isabelle experts. It explains how to define new
+logical systems, Isabelle's {\em raison d'\^etre}. Isabelle logics are
+hierarchies of theories. A number of simple examples are contained in {\em
+Introduction to Isabelle}; the full syntax of theory definition files ({\tt
+.thy} files) is shown in {\em The Isabelle Reference Manual}. This chapter's
+chief purpose is a thorough description of all syntax related matters
+concerning theories. The most important sections are \S\ref{sec:mixfix} about
+mixfix declarations and \S\ref{sec:macros} describing the macro system. The
+concluding examples of \S\ref{sec:min_logics} are more concerned with the
+logical aspects of the definition of theories. Sections marked with * can be
+skipped on the first reading.
+
+
+%% FIXME move to Refman
+% \section{Classes and types *}
+% \index{*arities!context conditions}
+%
+% Type declarations are subject to the following two well-formedness
+% conditions:
+% \begin{itemize}
+% \item There are no two declarations $ty :: (\vec{r})c$ and $ty :: (\vec{s})c$
+% with $\vec{r} \neq \vec{s}$. For example
+% \begin{ttbox}
+% types ty 1
+% arities ty :: ({\ttlbrace}logic{\ttrbrace}) logic
+% ty :: ({\ttlbrace}{\ttrbrace})logic
+% \end{ttbox}
+% leads to an error message and fails.
+% \item If there are two declarations $ty :: (s@1,\dots,s@n)c$ and $ty ::
+% (s@1',\dots,s@n')c'$ such that $c' < c$ then $s@i' \preceq s@i$ must hold
+% for $i=1,\dots,n$. The relationship $\preceq$, defined as
+% \[ s' \preceq s \iff \forall c\in s. \exists c'\in s'.~ c'\le c, \]
+% expresses that the set of types represented by $s'$ is a subset of the set of
+% types represented by $s$. For example
+% \begin{ttbox}
+% classes term < logic
+% types ty 1
+% arities ty :: ({\ttlbrace}logic{\ttrbrace})logic
+% ty :: ({\ttlbrace}{\ttrbrace})term
+% \end{ttbox}
+% leads to an error message and fails.
+% \end{itemize}
+% These conditions guarantee principal types~\cite{nipkow-prehofer}.
+
+
+
+\section{Precedence grammars} \label{sec:precedence_grammars}
+
+The precise syntax of a logic is best defined by a \rmindex{context-free
+grammar}. In order to simplify the description of mathematical languages, we
+introduce an extended format which permits {\bf
+precedences}\indexbold{precedence}. This scheme generalizes precedence
+declarations in \ML\ and {\sc prolog}. In this extended grammar format,
+nonterminals are decorated by integers, their precedence. In the sequel,
+precedences are shown as subscripts. A nonterminal $A@p$ on the right-hand
+side of a production may only be replaced using a production $A@q = \gamma$
+where $p \le q$.
+
+Formally, a set of context free productions $G$ induces a derivation
+relation $\rew@G$ on strings as follows:
+\[ \alpha A@p \beta ~\rew@G~ \alpha\gamma\beta ~~~iff~~~
+ \exists (A@q=\gamma) \in G.~q \ge p
+\]
+Any extended grammar of this kind can be translated into a normal context
+free grammar. However, this translation may require the introduction of a
+large number of new nonterminals and productions.
+
+\begin{example} \label{ex:precedence}
+The following simple grammar for arithmetic expressions demonstrates how
+binding power and associativity of operators can be enforced by precedences.
+\begin{center}
+\begin{tabular}{rclr}
+ $A@9$ & = & {\tt0} \\
+ $A@9$ & = & {\tt(} $A@0$ {\tt)} \\
+ $A@0$ & = & $A@0$ {\tt+} $A@1$ \\
+ $A@2$ & = & $A@3$ {\tt*} $A@2$ \\
+ $A@3$ & = & {\tt-} $A@3$
+\end{tabular}
+\end{center}
+The choice of precedences determines that {\tt -} binds tighter than {\tt *}
+which binds tighter than {\tt +}, and that {\tt +} associates to the left and
+{\tt *} to the right.
+\end{example}
+
+To minimize the number of subscripts, we adopt the following conventions:
+\begin{itemize}
+\item All precedences $p$ must be in the range $0 \leq p \leq max_pri$ for
+ some fixed $max_pri$.
+\item Precedence $0$ on the right-hand side and precedence $max_pri$ on the
+ left-hand side may be omitted.
+\end{itemize}
+In addition, we write the production $A@p = \alpha$ as $A = \alpha~(p)$,
+i.e.\ the precedence of the left-hand side actually appears in the uttermost
+right. Finally, alternatives may be separated by $|$, and repetition
+indicated by \dots.
+
+Using these conventions and assuming $max_pri=9$, the grammar in
+Example~\ref{ex:precedence} becomes
+\begin{center}
+\begin{tabular}{rclc}
+$A$ & = & {\tt0} & \hspace*{4em} \\
+ & $|$ & {\tt(} $A$ {\tt)} \\
+ & $|$ & $A$ {\tt+} $A@1$ & (0) \\
+ & $|$ & $A@3$ {\tt*} $A@2$ & (2) \\
+ & $|$ & {\tt-} $A@3$ & (3)
+\end{tabular}
+\end{center}
+
+
+
+\section{Basic syntax} \label{sec:basic_syntax}
+
+The basis of all extensions by object-logics is the \rmindex{Pure theory},
+bound to the \ML-identifier \ttindex{Pure.thy}. It contains, among many other
+things, the \rmindex{Pure syntax}. An informal account of this basic syntax
+(meta-logic, types etc.) may be found in {\em Introduction to Isabelle}. A
+more precise description using a precedence grammar is shown in
+Figure~\ref{fig:pure_gram}.
+
+\begin{figure}[htb]
+\begin{center}
+\begin{tabular}{rclc}
+$prop$ &=& \ttindex{PROP} $aprop$ ~~$|$~~ {\tt(} $prop$ {\tt)} \\
+ &$|$& $logic@3$ \ttindex{==} $logic@2$ & (2) \\
+ &$|$& $logic@3$ \ttindex{=?=} $logic@2$ & (2) \\
+ &$|$& $prop@2$ \ttindex{==>} $prop@1$ & (1) \\
+ &$|$& {\tt[|} $prop$ {\tt;} \dots {\tt;} $prop$ {\tt|]} {\tt==>} $prop@1$ & (1) \\
+ &$|$& {\tt!!} $idts$ {\tt.} $prop$ & (0) \\\\
+$logic$ &=& $prop$ ~~$|$~~ $fun$ \\\\
+$aprop$ &=& $id$ ~~$|$~~ $var$
+ ~~$|$~~ $fun@{max_pri}$ {\tt(} $logic$ {\tt,} \dots {\tt,} $logic$ {\tt)} \\\\
+$fun$ &=& $id$ ~~$|$~~ $var$ ~~$|$~~ {\tt(} $fun$ {\tt)} \\
+ &$|$& \ttindex{\%} $idts$ {\tt.} $logic$ & (0) \\\\
+$idts$ &=& $idt$ ~~$|$~~ $idt@1$ $idts$ \\\\
+$idt$ &=& $id$ ~~$|$~~ {\tt(} $idt$ {\tt)} \\
+ &$|$& $id$ \ttindex{::} $type$ & (0) \\\\
+$type$ &=& $tfree$ ~~$|$~~ $tvar$ \\
+ &$|$& $tfree$ {\tt::} $sort$ ~~$|$~~ $tvar$ {\tt::} $sort$ \\
+ &$|$& $id$ ~~$|$~~ $type@{max_pri}$ $id$
+ ~~$|$~~ {\tt(} $type$ {\tt,} \dots {\tt,} $type$ {\tt)} $id$ \\
+ &$|$& $type@1$ \ttindex{=>} $type$ & (0) \\
+ &$|$& {\tt[} $type$ {\tt,} \dots {\tt,} $type$ {\tt]} {\tt=>} $type$&(0)\\
+ &$|$& {\tt(} $type$ {\tt)} \\\\
+$sort$ &=& $id$ ~~$|$~~ {\tt\ttlbrace\ttrbrace}
+ ~~$|$~~ {\tt\ttlbrace} $id$ {\tt,} \dots {\tt,} $id$ {\tt\ttrbrace}
+\end{tabular}\index{*"!"!}\index{*"["|}\index{*"|"]}
+\indexbold{type@$type$} \indexbold{sort@$sort$} \indexbold{idt@$idt$}
+\indexbold{idts@$idts$} \indexbold{logic@$logic$} \indexbold{prop@$prop$}
+\indexbold{fun@$fun$}
+\end{center}
+\caption{Meta-Logic Syntax}
+\label{fig:pure_gram}
+\end{figure}
+
+The following main categories are defined:
+\begin{description}
+ \item[$prop$] Terms of type $prop$, i.e.\ formulae of the meta-logic.
+
+ \item[$aprop$] Atomic propositions.
+
+ \item[$logic$] Terms of types in class $logic$. Initially, $logic$ contains
+ merely $prop$. As the syntax is extended by new object-logics, more
+ productions for $logic$ are added automatically for (see below).
+
+ \item[$fun$] Terms potentially of function type.
+
+ \item[$type$] Meta-types.
+
+ \item[$idts$] a list of identifiers, possibly constrained by types. Note
+ that \verb|x :: nat y| is parsed as \verb|x :: (nat y)|, i.e.\ {\tt y}
+ would be treated like a type constructor applied to {\tt nat} here.
+\end{description}
+
+
+\subsection{Logical types and default syntax}
+
+Isabelle is concerned with mathematical languages which have a certain
+minimal vocabulary: identifiers, variables, parentheses, and the lambda
+calculus. Logical types, i.e.\ those of class $logic$, are automatically
+equipped with this basic syntax. More precisely, for any type constructor
+$ty$ with arity $(\dots)c$, where $c$ is a subclass of $logic$, the following
+productions are added:
+\begin{center}
+\begin{tabular}{rclc}
+$ty$ &=& $id$ ~~$|$~~ $var$ ~~$|$~~ {\tt(} $ty$ {\tt)} \\
+ &$|$& $fun@{max_pri}$ {\tt(} $logic$ {\tt,} \dots {\tt,} $logic$ {\tt)}\\
+ &$|$& $ty@{max_pri}$ {\tt::} $type$\\\\
+$logic$ &=& $ty$
+\end{tabular}
+\end{center}
+
+
+\subsection{Lexical matters *}
+
+The parser does not process input strings directly, it rather operates on
+token lists provided by Isabelle's \rmindex{lexical analyzer} (the
+\bfindex{lexer}). There are two different kinds of tokens: {\bf
+literals}\indexbold{literal token}\indexbold{token!literal} and {\bf valued
+tokens}\indexbold{valued token}\indexbold{token!valued}.
+
+Literals (or delimiters \index{delimiter}) can be regarded as reserved
+words\index{reserved word} of the syntax and the user can add new ones, when
+extending theories. In Figure~\ref{fig:pure_gram} they appear in typewriter
+type, e.g.\ {\tt PROP}, {\tt ==}, {\tt =?=}, {\tt ;}.
+
+Valued tokens on the other hand have a fixed predefined syntax. The lexer
+distinguishes four kinds of them: identifiers\index{identifier},
+unknowns\index{unknown}\index{scheme variable|see{unknown}}, type
+variables\index{type variable}, type unknowns\index{type unknown}\index{type
+scheme variable|see{type unknown}}; they are denoted by $id$\index{id@$id$},
+$var$\index{var@$var$}, $tfree$\index{tfree@$tfree$},
+$tvar$\index{tvar@$tvar$}, respectively. Typical examples are {\tt x}, {\tt
+?x7}, {\tt 'a}, {\tt ?'a3}, the exact syntax is:
+
+\begin{tabular}{rcl}
+$id$ & = & $letter~quasiletter^*$ \\
+$var$ & = & ${\tt ?}id ~~|~~ {\tt ?}id{\tt .}nat$ \\
+$tfree$ & = & ${\tt '}id$ \\
+$tvar$ & = & ${\tt ?}tfree ~~|~~ {\tt ?}tfree{\tt .}nat$ \\[1ex]
+
+$letter$ & = & one of {\tt a}\dots {\tt z} {\tt A}\dots {\tt Z} \\
+$digit$ & = & one of {\tt 0}\dots {\tt 9} \\
+$quasiletter$ & = & $letter ~~|~~ digit ~~|~~ {\tt _} ~~|~~ {\tt '}$ \\
+$nat$ & = & $digit^+$
+\end{tabular}
+
+A string of $var$ or $tvar$ describes an \rmindex{unknown} which is
+internally a pair of base name and index (\ML\ type \ttindex{indexname}).
+These components are either explicitly separated by a dot as in {\tt ?x.1} or
+{\tt ?x7.3} or directly run together as in {\tt ?x1}. The latter form is
+possible, if the base name does not end with digits. Additionally, if the
+index is 0, it may be dropped altogether, e.g.\ {\tt ?x} abbreviates {\tt
+?x0} or {\tt ?x.0}.
+
+Note that $id$, $var$, $tfree$, $tvar$ are mutually disjoint, but it is
+perfectly legal that they overlap with the set of literal tokens (e.g.\ {\tt
+PROP}, {\tt ALL}, {\tt EX}).
+
+The lexical analyzer translates input strings to token lists by repeatedly
+taking the maximal prefix of the input string that forms a valid token. A
+maximal prefix that is both a literal and a valued token is treated as a
+literal. Spaces, tabs and newlines are separators; they never occur within
+tokens.
+
+Note that literals need not necessarily be surrounded by white space to be
+recognized as separate. For example, if {\tt -} is a literal but {\tt --} is
+not, then the string {\tt --} is treated as two consecutive occurrences of
+{\tt -}. This is in contrast to \ML\ which would treat {\tt --} always as a
+single symbolic name. The consequence of Isabelle's more liberal scheme is
+that the same string may be parsed in different ways after extending the
+syntax: after adding {\tt --} as a literal, the input {\tt --} is treated as
+a single token.
+
+
+\subsection{Inspecting syntax *}
+
+You may get the \ML\ representation of the syntax of any Isabelle theory by
+applying \index{*syn_of}
+\begin{ttbox}
+ syn_of: theory -> Syntax.syntax
+\end{ttbox}
+\ttindex{Syntax.syntax} is an abstract type containing quite a lot of
+material, wherefore there is no toplevel pretty printer set up for displaying
+it automatically. You have to call one of the following functions instead:
+\index{*Syntax.print_syntax} \index{*Syntax.print_gram}
+\index{*Syntax.print_trans}
+\begin{ttbox}
+ Syntax.print_syntax: Syntax.syntax -> unit
+ Syntax.print_gram: Syntax.syntax -> unit
+ Syntax.print_trans: Syntax.syntax -> unit
+\end{ttbox}
+where {\tt Syntax.print_syntax} shows virtually all information contained in
+a syntax, therefore being quite verbose. Its output is divided into labeled
+sections, the syntax proper is represented by {\tt lexicon}, {\tt roots} and
+{\tt prods}. The rest refers to the manifold facilities to apply syntactic
+translations to parse trees (macro expansion etc.). See \S\ref{sec:macros}
+and \S\ref{sec:tr_funs} for more details on translations.
+
+To simplify coping with the verbosity of {\tt Syntax.print_syntax}, there are
+\ttindex{Syntax.print_gram} to print the syntax proper only and
+\ttindex{Syntax.print_trans} to print the translation related stuff only.
+
+Let's have a closer look at part of Pure's syntax:
+\begin{ttbox}
+Syntax.print_syntax (syn_of Pure.thy);
+{\out lexicon: "!!" "%" "(" ")" "," "." "::" ";" "==" "==>" \dots}
+{\out roots: logic type fun prop}
+{\out prods:}
+{\out type = tfree (1000)}
+{\out type = tvar (1000)}
+{\out type = id (1000)}
+{\out type = tfree "::" sort[0] => "_ofsort" (1000)}
+{\out type = tvar "::" sort[0] => "_ofsort" (1000)}
+{\out \vdots}
+{\out consts: "_K" "_appl" "_aprop" "_args" "_asms" "_bigimpl" \dots}
+{\out parse_ast_translation: "_appl" "_bigimpl" "_bracket"}
+{\out "_idtyp" "_lambda" "_tapp" "_tappl"}
+{\out parse_rules:}
+{\out parse_translation: "!!" "_K" "_abs" "_aprop"}
+{\out print_translation: "all"}
+{\out print_rules:}
+{\out print_ast_translation: "==>" "_abs" "_idts" "fun"}
+\end{ttbox}
+
+\begin{description}
+ \item[\ttindex{lexicon}]
+ The set of literal tokens (i.e.\ reserved words, delimiters) used for
+ lexical analysis.
+
+ \item[\ttindex{roots}]
+ The legal syntactic categories to start parsing with. You name the
+ desired root directly as a string when calling lower level functions or
+ specifying macros. Higher level functions usually expect a type and
+ derive the actual root as follows:\index{root_of_type@$root_of_type$}
+ \begin{itemize}
+ \item $root_of_type(\tau@1 \To \tau@2) = \mtt{fun}$.
+
+ \item $root_of_type(\tau@1 \mathrel{ty} \tau@2) = ty$.
+
+ \item $root_of_type((\tau@1, \dots, \tau@n)ty) = ty$.
+
+ \item $root_of_type(\alpha) = \mtt{logic}$.
+ \end{itemize}
+ Thereby $\tau@1, \dots, \tau@n$ are types, $ty$ a type infix or ordinary
+ type constructor and $\alpha$ a type variable or unknown. Note that only
+ the outermost type constructor is taken into account.
+
+ \item[\ttindex{prods}]
+ The list of productions describing the precedence grammar. Nonterminals
+ $A@n$ are rendered in ASCII as {\tt $A$[$n$]}, literal tokens are quoted.
+ Note that some productions have strings attached after an {\tt =>}. These
+ strings later become the heads of parse trees, but they also play a vital
+ role when terms are printed (see \S\ref{sec:asts}).
+
+ Productions which do not have string attached and thus do not create a
+ new parse tree node are called {\bf copy productions}\indexbold{copy
+ production}. They must have exactly one
+ argument\index{argument!production} (i.e.\ nonterminal or valued token)
+ on the right-hand side. The parse tree generated when parsing that
+ argument is simply passed up as the result of parsing the whole copy
+ production.
+
+ A special kind of copy production is one where the argument is a
+ nonterminal and no additional syntactic sugar (literals) exists. It is
+ called a \bfindex{chain production}. Chain productions should be seen as
+ an abbreviation mechanism: conceptually, they are removed from the
+ grammar by adding appropriate new rules. Precedence information attached
+ to chain productions is ignored, only the dummy value $-1$ is displayed.
+
+ \item[\ttindex{consts}, \ttindex{parse_rules}, \ttindex{print_rules}]
+ This is macro related stuff (see \S\ref{sec:macros}).
+
+ \item[\tt *_translation]
+ \index{*parse_ast_translation} \index{*parse_translation}
+ \index{*print_translation} \index{*print_ast_translation}
+ The sets of constants that invoke translation functions. These are more
+ arcane matters (see \S\ref{sec:asts} and \S\ref{sec:tr_funs}).
+\end{description}
+
+Of course you may inspect the syntax of any theory using the above calling
+sequence. But unfortunately, as more and more material accumulates, the
+output becomes even more verbose. Note that when extending syntaxes new {\tt
+roots}, {\tt prods}, {\tt parse_rules} and {\tt print_rules} are appended to
+the end. The other lists are displayed sorted.
+
+
+
+\section{Abstract syntax trees} \label{sec:asts}
+
+Figure~\ref{fig:parse_print} shows a simplified model of the parsing and
+printing process.
+
+\begin{figure}[htb]
+\begin{center}
+\begin{tabular}{cl}
+string & \\
+$\downarrow$ & parser \\
+parse tree & \\
+$\downarrow$ & \rmindex{parse ast translation} \\
+ast & \\
+$\downarrow$ & ast rewriting (macros) \\
+ast & \\
+$\downarrow$ & \rmindex{parse translation}, type-inference \\
+--- well-typed term --- & \\
+$\downarrow$ & \rmindex{print translation} \\
+ast & \\
+$\downarrow$ & ast rewriting (macros) \\
+ast & \\
+$\downarrow$ & \rmindex{print ast translation}, printer \\
+string &
+\end{tabular}
+\end{center}
+\caption{Parsing and Printing}
+\label{fig:parse_print}
+\end{figure}
+
+The parser takes an input string (well, actually a token list produced by the
+lexer) and produces a \rmindex{parse tree}, which is directly derived from
+the productions involved. By applying some internal transformations the parse
+tree becomes an \bfindex{abstract syntax tree} (\bfindex{ast}), macro
+expansion, further translations and finally type inference yields a
+well-typed term\index{term!well-typed}.
+
+The printing process is the reverse, except for some subtleties to be
+discussed later.
+
+Asts are an intermediate form between the very raw parse trees and the typed
+$\lambda$-terms. An ast is either an atom (constant or variable) or a list of
+{\em at least two} subasts. Internally, they have type \ttindex{Syntax.ast}
+which is defined as: \index{*Constant} \index{*Variable} \index{*Appl}
+\begin{ttbox}
+datatype ast =
+ Constant of string |
+ Variable of string |
+ Appl of ast list
+\end{ttbox}
+
+Notation: We write constant atoms as quoted strings, variable atoms as
+non-quoted strings and applications as list of subasts separated by space and
+enclosed in parentheses. For example:
+\begin{ttbox}
+ Appl [Constant "_constrain",
+ Appl [Constant "_abs", Variable "x", Variable "t"],
+ Appl [Constant "fun", Variable "'a", Variable "'b"]]
+ {\rm is written as} ("_constrain" ("_abs" x t) ("fun" 'a 'b))
+\end{ttbox}
+
+Note that {\tt ()} and {\tt (f)} are both illegal.
+
+The resemblance of LISP's S-expressions is intentional, but notice the two
+kinds of atomic symbols: $\Constant x$ and $\Variable x$. This distinction
+has some more obscure reasons and you can ignore it at about half of the
+time. By now you shouldn't take the names ``{\tt Constant}'' and ``{\tt
+Variable}'' too literally. In the later translation to terms $\Variable x$
+may become a constant, free or bound variable, even a type constructor or
+class name; the actual outcome depends on the context.
+
+Similarly, you can think of ${\tt (} f~x@1~\ldots~x@n{\tt )}$ as some sort of
+application of $f$ to the arguments $x@1, \ldots, x@n$. But the actual kind
+of application (say a type constructor $f$ applied to types $x@1, \ldots,
+x@n$) is determined later by context, too.
+
+Forms like {\tt (("_abs" x $t$) $u$)} are perfectly legal, but asts are not
+higher-order: the {\tt "_abs"} does not yet bind the {\tt x} in any way,
+though later at the term level, {\tt ("_abs" x $t$)} will become an {\tt Abs}
+node and occurrences of {\tt x} in $t$ will be replaced by {\tt Bound}s. Even
+if non constant heads of applications may seem unusual, asts should be
+regarded as first order. Purists may think of ${\tt (} f~x@1~\ldots~x@n{\tt
+)}$ as a first order application of some invisible $(n+1)$-ary constant.
+
+
+\subsection{Parse trees to asts}
+
+Asts are generated from parse trees by applying some translation functions,
+which are internally called {\bf parse ast translations}\indexbold{parse ast
+translation}.
+
+We can think of the raw output of the parser as trees built up by nesting the
+right-hand sides of those productions that were used to derive a word that
+matches the input token list. Then parse trees are simply lists of tokens and
+sub parse trees, the latter replacing the nonterminals of the productions.
+Note that we refer here to the actual productions in their internal form as
+displayed by {\tt Syntax.print_syntax}.
+
+Ignoring parse ast translations, the mapping
+$ast_of_pt$\index{ast_of_pt@$ast_of_pt$} from parse trees to asts is derived
+from the productions involved as follows:
+\begin{itemize}
+ \item Valued tokens: $ast_of_pt(t) = \Variable s$, where $t$ is an $id$,
+ $var$, $tfree$ or $tvar$ token, and $s$ its value.
+
+ \item Copy productions: $ast_of_pt(\ldots P \ldots) = ast_of_pt(P)$.
+
+ \item $0$-ary productions: $ast_of_pt(\ldots \mtt{=>} c) = \Constant c$.
+
+ \item $n$-ary productions: $ast_of_pt(\ldots P@1 \ldots P@n \ldots \mtt{=>}
+ c) = \Appl{\Constant c,$ $ast_of_pt(P@1),$ $\ldots,$ $ast_of_pt(P@n)}$,
+ where $n \ge 1$.
+\end{itemize}
+Thereby $P, P@1, \ldots, P@n$ denote sub parse trees or valued tokens and
+``\dots'' zero or more literal tokens. That means literal tokens are stripped
+and do not appear in asts.
+
+The following table presents some simple examples:
+
+{\tt\begin{tabular}{ll}
+\rm input string & \rm ast \\\hline
+"f" & f \\
+"'a" & 'a \\
+"t == u" & ("==" t u) \\
+"f(x)" & ("_appl" f x) \\
+"f(x, y)" & ("_appl" f ("_args" x y)) \\
+"f(x, y, z)" & ("_appl" f ("_args" x ("_args" y z))) \\
+"\%x y.\ t" & ("_lambda" ("_idts" x y) t) \\
+\end{tabular}}
+
+Some of these examples illustrate why further translations are desirable in
+order to provide some nice standard form of asts before macro expansion takes
+place. Hence the Pure syntax provides predefined parse ast
+translations\index{parse ast translation!of Pure} for ordinary applications,
+type applications, nested abstraction, meta implication and function types.
+Their net effect on some representative input strings is shown in
+Figure~\ref{fig:parse_ast_tr}.
+
+\begin{figure}[htb]
+\begin{center}
+{\tt\begin{tabular}{ll}
+\rm string & \rm ast \\\hline
+"f(x, y, z)" & (f x y z) \\
+"'a ty" & (ty 'a) \\
+"('a, 'b) ty" & (ty 'a 'b) \\
+"\%x y z.\ t" & ("_abs" x ("_abs" y ("_abs" z t))) \\
+"\%x ::\ 'a.\ t" & ("_abs" ("_constrain" x 'a) t) \\
+"[| P; Q; R |] => S" & ("==>" P ("==>" Q ("==>" R S))) \\
+"['a, 'b, 'c] => 'd" & ("fun" 'a ("fun" 'b ("fun" 'c 'd)))
+\end{tabular}}
+\end{center}
+\caption{Built-in Parse Ast Translations}
+\label{fig:parse_ast_tr}
+\end{figure}
+
+This translation process is guided by names of constant heads of asts. The
+list of constants invoking parse ast translations is shown in the output of
+{\tt Syntax.print_syntax} under {\tt parse_ast_translation}.
+
+
+\subsection{Asts to terms *}
+
+After the ast has been normalized by the macro expander (see
+\S\ref{sec:macros}), it is transformed into a term with yet another set of
+translation functions interspersed: {\bf parse translations}\indexbold{parse
+translation}. The result is a non-typed term\index{term!non-typed} which may
+contain type constraints, that is 2-place applications with head {\tt
+"_constrain"}. The second argument of a constraint is a type encoded as term.
+Real types will be introduced later during type inference, which is not
+discussed here.
+
+If we ignore the built-in parse translations of Pure first, then the mapping
+$term_of_ast$\index{term_of_ast@$term_of_ast$} from asts to (non-typed) terms
+is defined by:
+\begin{itemize}
+ \item $term_of_ast(\Constant x) = \ttfct{Const} (x, \mtt{dummyT})$.
+
+ \item $term_of_ast(\Variable \mtt{"?}xi\mtt") = \ttfct{Var} ((x, i),
+ \mtt{dummyT})$, where $x$ is the base name and $i$ the index extracted
+ from $xi$.
+
+ \item $term_of_ast(\Variable x) = \ttfct{Free} (x, \mtt{dummyT})$.
+
+ \item $term_of_ast(\Appl{f, x@1, \ldots, x@n}) = term_of_ast(f) ~{\tt\$}~
+ term_of_ast(x@1)$ {\tt\$} \dots\ {\tt\$} $term_of_ast(x@n)$.
+\end{itemize}
+Note that {\tt Const}, {\tt Var}, {\tt Free} belong to the datatype {\tt
+term} and \ttindex{dummyT} is bound to some dummy {\tt typ}, which is ignored
+during type-inference.
+
+So far the outcome is still a first order term. {\tt Abs}s and {\tt Bound}s
+are introduced by parse translations associated with {\tt "_abs"} and {\tt
+"!!"} (and any other user defined binder).
+
+
+\subsection{Printing of terms *}
+
+When terms are prepared for printing, they are first transformed into asts
+via $ast_of_term$\index{ast_of_term@$ast_of_term$} (ignoring {\bf print
+translations}\indexbold{print translation}):
+\begin{itemize}
+ \item $ast_of_term(\ttfct{Const} (x, \tau)) = \Constant x$.
+
+ \item $ast_of_term(\ttfct{Free} (x, \tau)) = constrain (\Variable x,
+ \tau)$.
+
+ \item $ast_of_term(\ttfct{Var} ((x, i), \tau)) = constrain (\Variable
+ \mtt{"?}xi\mtt", \tau)$, where $\mtt?xi$ is the string representation of
+ the {\tt indexname} $(x, i)$.
+
+ \item $ast_of_term(\ttfct{Abs} (x, \tau, t)) = \ttfct{Appl}
+ \mathopen{\mtt[} \Constant \mtt{"_abs"}, constrain(\Variable x', \tau),$
+ $ast_of_term(t') \mathclose{\mtt]}$, where $x'$ is a version of $x$
+ renamed away from all names occurring in $t$, and $t'$ the body $t$ with
+ all {\tt Bound}s referring to this {\tt Abs} replaced by $\ttfct{Free}
+ (x', \mtt{dummyT})$.
+
+ \item $ast_of_term(\ttfct{Bound} i) = \Variable \mtt{"B.}i\mtt"$. Note that
+ the occurrence of loose bound variables is abnormal and should never
+ happen when printing well-typed terms.
+
+ \item $ast_of_term(f \ttrel{\$} x@1 \ttrel{\$} \ldots \ttrel{\$} x@n) =
+ \ttfct{Appl} \mathopen{\mtt[} ast_of_term(f), ast_of_term(x@1), \ldots,$
+ $ast_of_term(x@n) \mathclose{\mtt]}$, where $f$ is not an application.
+
+ \item $constrain(x, \tau) = x$, if $\tau = \mtt{dummyT}$ \index{*dummyT} or
+ \ttindex{show_types} not set to {\tt true}.
+
+ \item $constrain(x, \tau) = \Appl{\Constant \mtt{"_constrain"}, x, ty}$,
+ where $ty$ is the ast encoding of $\tau$. That is: type constructors as
+ {\tt Constant}s, type variables as {\tt Variable}s and type applications
+ as {\tt Appl}s with the head type constructor as first element.
+ Additionally, if \ttindex{show_sorts} is set to {\tt true}, some type
+ variables are decorated with an ast encoding their sort.
+\end{itemize}
+
+\medskip
+After an ast has been normalized wrt.\ the print macros, it is transformed
+into the final output string. The built-in {\bf print ast
+translations}\indexbold{print ast translation} are essentially the reverse
+ones of the parse ast translations of Figure~\ref{fig:parse_ast_tr}.
+
+For the actual printing process, the names attached to grammar productions of
+the form $\ldots A@{p@1}^1 \ldots A@{p@n}^n \ldots \mtt{=>} c$ play a vital
+role. Whenever an ast with constant head $c$, i.e.\ $\mtt"c\mtt"$ or
+$(\mtt"c\mtt"~ x@1 \ldots x@n)$, is encountered it is printed according to
+the production for $c$. This means that first the arguments $x@1$ \dots $x@n$
+are printed, then put in parentheses if necessary for reasons of precedence,
+and finally joined to a single string, separated by the syntactic sugar of
+the production (denoted by ``\dots'' above).
+
+If an application $(\mtt"c\mtt"~ x@1 \ldots x@m)$ has more arguments than the
+corresponding production, it is first split into $((\mtt"c\mtt"~ x@1 \ldots
+x@n)~ (x@{n+1} \ldots x@m))$. Applications with too few arguments or with
+non-constant head or without a corresponding production are printed as
+$f(x@1, \ldots, x@l)$ or $(\alpha@1, \ldots, \alpha@l) ty$. A single
+$\Variable x$ is simply printed as $x$.
+
+Note that the system does {\em not} insert blanks automatically. They should
+be part of the mixfix declaration (which provide the user interface for
+defining syntax) if they are required to separate tokens. Mixfix declarations
+may also contain pretty printing annotations. See \S\ref{sec:mixfix} for
+details.
+
+
+
+\section{Mixfix declarations} \label{sec:mixfix}
+
+When defining theories, the user usually declares new constants as names
+associated with a type followed by an optional \bfindex{mixfix annotation}.
+If none of the constants are introduced with mixfix annotations, there is no
+concrete syntax to speak of: terms can only be abstractions or applications
+of the form $f(t@1, \dots, t@n)$. Since this notation quickly becomes
+unreadable, Isabelle supports syntax definitions in the form of unrestricted
+context-free \index{context-free grammar} \index{precedence grammar}
+precedence grammars using mixfix annotations.
+
+Mixfix annotations describe the {\em concrete} syntax, a standard translation
+into the abstract syntax and a pretty printing scheme, all in one. Isabelle
+syntax definitions are inspired by \OBJ's~\cite{OBJ} {\em mixfix\/} syntax.
+Each mixfix annotation defines a precedence grammar production and optionally
+associates a constant with it.
+
+There is a general form of mixfix annotation exhibiting the full power of
+extending a theory's syntax, and some shortcuts for common cases like infix
+operators.
+
+The general \bfindex{mixfix declaration} as it may occur within a {\tt
+consts} section\index{consts section@{\tt consts} section} of a {\tt .thy}
+file, specifies a constant declaration and a grammar production at the same
+time. It has the form {\tt $c$ ::\ "$\tau$" ("$sy$" $ps$ $p$)} and is
+interpreted as follows:
+\begin{itemize}
+ \item {\tt $c$ ::\ "$\tau$"} is the actual constant declaration without any
+ syntactic significance.
+
+ \item $sy$ is the right-hand side of the production, specified as a
+ symbolic string. In general, $sy$ is of the form $\alpha@0 \_ \alpha@1
+ \dots \alpha@{n-1} \_ \alpha@n$, where each occurrence of \ttindex{_}
+ denotes an argument\index{argument!mixfix} position and the strings
+ $\alpha@i$ do not contain (non-escaped) {\tt _}. The $\alpha@i$ may
+ consist of delimiters\index{delimiter},
+ spaces\index{space (pretty printing)} or \rmindex{pretty printing}
+ annotations (see below).
+
+ \item $\tau$ specifies the syntactic categories of the arguments on the
+ left-hand and of the right-hand side. Arguments may be nonterminals or
+ valued tokens. If $sy$ is of the form above, $\tau$ must be a nested
+ function type with at least $n$ argument positions, say $\tau = [\tau@1,
+ \dots, \tau@n] \To \tau'$. The syntactic category of argument $i$ is
+ derived from type $\tau@i$ (see $root_of_type$ in
+ \S\ref{sec:basic_syntax}). The result, i.e.\ the left-hand side of the
+ production, is derived from type $\tau'$. Note that the $\tau@i$ and
+ $\tau'$ may be function types.
+
+ \item $c$ is the name of the constant associated with the production. If
+ $c$ is the empty string {\tt ""} then this is a \rmindex{copy
+ production}. Otherwise, parsing an instance of the phrase $sy$ generates
+ the ast {\tt ("$c$" $a@1$ $\ldots$ $a@n$)}, where $a@i$ is the ast
+ generated by parsing the $i$-th argument.
+
+ \item $ps$ is an optional list of at most $n$ integers, say {\tt [$p@1$,
+ $\ldots$, $p@m$]}, where $p@i$ is the minimal \rmindex{precedence}
+ required of any phrase that may appear as the $i$-th argument. Missing
+ precedences default to $0$.
+
+ \item $p$ is the \rmindex{precedence} the of this production.
+\end{itemize}
+
+Precedences\index{precedence!range of} may range between $0$ and
+$max_pri$\indexbold{max_pri@$max_pri$} (= 1000). If you want to ignore them,
+the safest way to do so is to use the declaration {\tt $c$ ::\ "$\tau$"
+("$sy$")}. The resulting production puts no precedence constraints on any of
+its arguments and has maximal precedence itself.
+
+\begin{example}
+The following theory specification contains a {\tt consts} section that
+encodes the precedence grammar of Example~\ref{ex:precedence} as mixfix
+declarations:
+\begin{ttbox}
+EXP = Pure +
+types
+ exp 0
+arities
+ exp :: logic
+consts
+ "0" :: "exp" ("0" 9)
+ "+" :: "[exp, exp] => exp" ("_ + _" [0, 1] 0)
+ "*" :: "[exp, exp] => exp" ("_ * _" [3, 2] 2)
+ "-" :: "exp => exp" ("- _" [3] 3)
+end
+\end{ttbox}
+Note that the {\tt arities} declaration causes {\tt exp} to be added to the
+syntax' roots. If you put the above text into a file {\tt exp.thy} and load
+it via {\tt use_thy "exp"}, you can do some tests:
+\begin{ttbox}
+val read_exp = Syntax.test_read (syn_of EXP.thy) "exp";
+read_exp "0 * 0 * 0 * 0 + 0 + 0 + 0";
+{\out tokens: "0" "*" "0" "*" "0" "*" "0" "+" "0" "+" "0" "+" "0"}
+{\out raw: ("+" ("+" ("+" ("*" "0" ("*" "0" ("*" "0" "0"))) "0") "0") "0")}
+{\out \vdots}
+read_exp "0 + - 0 + 0";
+{\out tokens: "0" "+" "-" "0" "+" "0"}
+{\out raw: ("+" ("+" "0" ("-" "0")) "0")}
+{\out \vdots}
+\end{ttbox}
+The output of \ttindex{Syntax.test_read} includes the token list ({\tt
+tokens}) and the raw ast directly derived from the parse tree, ignoring parse
+ast translations. The rest is tracing information provided by the macro
+expander (see \S\ref{sec:macros}).
+
+Issuing {\tt Syntax.print_gram (syn_of EXP.thy)} will reveal the actual
+grammar productions derived from the above mixfix declarations (lots of
+additional stuff deleted):
+\begin{ttbox}
+exp = "0" => "0" (9)
+exp = exp[0] "+" exp[1] => "+" (0)
+exp = exp[3] "*" exp[2] => "*" (2)
+exp = "-" exp[3] => "-" (3)
+\end{ttbox}
+\end{example}
+
+Let us now have a closer look at the structure of the string $sy$ appearing
+in mixfix annotations. This string specifies a list of parsing and printing
+directives, namely delimiters\index{delimiter},
+arguments\index{argument!mixfix}, spaces\index{space (pretty printing)},
+blocks of indentation\index{block (pretty printing)} and optional or forced
+line breaks\index{break (pretty printing)}. These are encoded via the
+following character sequences:
+
+\begin{description}
+ \item[~\ttindex_~] An argument\index{argument!mixfix} position.
+
+ \item[~$d$~] A \rmindex{delimiter}, i.e.\ a non-empty sequence of
+ non-special or escaped characters. Escaping a
+ character\index{escape character} means preceding it with a {\tt '}
+ (quote). Thus you have to write {\tt ''} if you really want a single
+ quote. Delimiters may never contain white space, though.
+
+ \item[~$s$~] A non-empty sequence of spaces\index{space (pretty printing)}
+ for printing.
+
+ \item[~{\ttindex($n$}~] Open a block\index{block (pretty printing)}. $n$ is
+ an optional sequence of digits that specifies the amount of indentation
+ to be added when a line break occurs within the block. If {\tt(} is not
+ followed by a digit, the indentation defaults to $0$.
+
+ \item[~\ttindex)~] Close a block.
+
+ \item[~\ttindex{//}~] Force a line break\index{break (pretty printing)}.
+
+ \item[~\ttindex/$s$~] Allow a line break\index{break (pretty printing)}.
+ Spaces $s$ right after {\tt /} are only printed if the break is not
+ taken.
+\end{description}
+
+In terms of parsing, arguments are nonterminals or valued tokens, while
+delimiters are literal tokens. The other directives have only significance
+for printing. The \rmindex{pretty printing} mechanisms of Isabelle is
+essentially that described in \cite{FIXME:ML_for_the_Working_Programmer}.
+
+
+\subsection{Infixes}
+
+Infix\index{infix} operators associating to the left or right can be declared
+conveniently using \ttindex{infixl} or \ttindex{infixr}.
+
+Roughly speaking, the form {\tt $c$ ::\ "$\tau$" (infixl $p$)} abbreviates:
+\begin{ttbox}
+"op \(c\)" ::\ "\(\tau\)" ("op \(c\)")
+"op \(c\)" ::\ "\(\tau\)" ("(_ \(c\)/ _)" [\(p\), \(p + 1\)] \(p\))
+\end{ttbox}
+and {\tt $c$ ::\ "$\tau$" (infixr $p$)} abbreviates:
+\begin{ttbox}
+"op \(c\)" ::\ "\(\tau\)" ("op \(c\)")
+"op \(c\)" ::\ "\(\tau\)" ("(_ \(c\)/ _)" [\(p + 1\), \(p\)] \(p\))
+\end{ttbox}
+
+Thus, prefixing infixes with \ttindex{op} makes them behave like ordinary
+function symbols. Special characters occurring in $c$ have to be escaped as
+in delimiters. Also note that the expanded forms above would be actually
+illegal at the user level because of duplicate declarations of constants.
+
+
+\subsection{Binders}
+
+A \bfindex{binder} is a variable-binding construct, such as a
+\rmindex{quantifier}. The constant declaration \indexbold{*binder}
+\begin{ttbox}
+\(c\) ::\ "\(\tau\)" (binder "\(Q\)" \(p\))
+\end{ttbox}
+introduces a binder $c$ of type $\tau$, which must have the form $(\tau@1 \To
+\tau@2) \To \tau@3$. Its concrete syntax is $Q~x.P$. A binder is like a
+generalized quantifier where $\tau@1$ is the type of the bound variable $x$,
+$\tau@2$ the type of the body $P$, and $\tau@3$ the type of the whole term.
+For example $\forall$ can be declared like this:
+\begin{ttbox}
+All :: "('a => o) => o" (binder "ALL " 10)
+\end{ttbox}
+This allows us to write $\forall x.P$ either as {\tt All(\%$x$.$P$)} or {\tt
+ALL $x$.$P$}. When printing terms, Isabelle usually uses the latter form, but
+has to fall back on $\mtt{All}(P)$, if $P$ is not an abstraction.
+
+Binders $c$ of type $(\sigma \To \tau) \To \tau$ can be nested; then the
+internal form $c(\lambda x@1. c(\lambda x@2. \ldots c(\lambda x@n. P)
+\ldots))$ corresponds to external $Q~x@1~x@2 \ldots x@n. P$.
+
+\medskip
+The general binder declaration
+\begin{ttbox}
+\(c\) ::\ "(\(\tau@1\) => \(\tau@2\)) => \(\tau@3\)" (binder "\(Q\)" \(p\))
+\end{ttbox}
+is internally expanded to
+\begin{ttbox}
+\(c\) ::\ "(\(\tau@1\) => \(\tau@2\)) => \(\tau@3\)"
+"\(Q\)" ::\ "[idts, \(\tau@2\)] => \(\tau@3\)" ("(3\(Q\)_./ _)" \(p\))
+\end{ttbox}
+with $idts$ being the syntactic category for a list of $id$s optionally
+constrained (see Figure~\ref{fig:pure_gram}). Note that special characters in
+$Q$ have to be escaped as in delimiters.
+
+Additionally, a parse translation\index{parse translation!for binder} for $Q$
+and a print translation\index{print translation!for binder} for $c$ is
+installed. These perform behind the scenes the translation between the
+internal and external forms.
+
+
+
+\section{Syntactic translations (macros)} \label{sec:macros}
+
+So far we have pretended that there is a close enough relationship between
+concrete and abstract syntax to allow an automatic translation from one to
+the other using the constant name supplied with non-copy production. In many
+cases this scheme is not powerful enough. Some typical examples involve
+variable binding constructs (e.g.\ {\tt ALL x:A.P} vs.\ {\tt Ball(A, \%x.P)}
+or convenient notations for enumeration like finite sets, lists etc.\ (e.g.\
+{\tt [x, y, z]} vs.\ {\tt Cons(x, Cons(y, Cons(z, Nil)))}).
+
+Isabelle offers such translation facilities at two different levels, namely
+{\bf macros}\indexbold{macro} and {\bf translation functions}.
+
+Macros are specified by first order rewriting systems that operate on asts.
+They are usually easy to read and in most cases not very difficult to write.
+But some more obscure translations cannot be expressed as macros and you have
+to fall back on the more powerful mechanism of translation functions written
+in \ML. These are quite hard to write and nearly unreadable by the casual
+user (see \S\ref{sec:tr_funs}).
+
+\medskip
+Let us now getting started with the macro system by a simple example:
+
+\begin{example}~ \label{ex:set_trans}
+
+\begin{ttbox}
+SET = Pure +
+types
+ i, o 0
+arities
+ i, o :: logic
+consts
+ Trueprop :: "o => prop" ("_" 5)
+ Collect :: "[i, i => o] => i"
+ "{\at}Collect" :: "[idt, i, o] => i" ("(1{\ttlbrace}_:_./ _{\ttrbrace})")
+ Replace :: "[i, [i, i] => o] => i"
+ "{\at}Replace" :: "[idt, idt, i, o] => i" ("(1{\ttlbrace}_./ _:_, _{\ttrbrace})")
+ Ball :: "[i, i => o] => o"
+ "{\at}Ball" :: "[idt, i, o] => o" ("(3ALL _:_./ _)" 10)
+translations
+ "{\ttlbrace}x:A. P{\ttrbrace}" == "Collect(A, %x. P)"
+ "{\ttlbrace}y. x:A, Q{\ttrbrace}" == "Replace(A, %x y. Q)"
+ "ALL x:A. P" == "Ball(A, %x. P)"
+end
+\end{ttbox}
+
+This and the following theories are complete working examples, though they
+are fragmentary as they contain merely syntax. They are somewhat fashioned
+after {\tt ZF/zf.thy}, where you should look for a good real-world example.
+
+{\tt SET} defines constants for set comprehension ({\tt Collect}),
+replacement ({\tt Replace}) and bounded universal quantification ({\tt
+Ball}). Without additional syntax you would have to express $\forall x \in A.
+P(x)$ as {\tt Ball(A, P)}. Since this is quite awkward, we define additional
+constants with appropriate concrete syntax. These constants are decorated
+with {\tt\at} to stress their pure syntactic purpose; they should never occur
+within the final well-typed terms. Another consequence is that the user
+cannot 'forge' terms containing such names, like {\tt\at Collect(x)}, for
+being syntactically incorrect.
+
+The included translations cause the replacement of external forms by internal
+forms after parsing and vice versa before printing of terms.
+\end{example}
+
+This is only a very simple but common instance of a more powerful mechanism.
+As a specification of what is to be translated, it should be comprehensible
+without further explanations. But there are also some snags and other
+peculiarities that are typical for macro systems in general. The purpose of
+this section is to explain how Isabelle's macro system really works.
+
+
+\subsection{Specifying macros}
+
+Basically macros are rewrite rules on asts. But unlike other macro systems of
+various programming languages, Isabelle's macros work two way. Therefore a
+syntax contains two lists of rules: one for parsing and one for printing.
+
+The {\tt translations} section\index{translations section@{\tt translations}
+section} consists of a list of rule specifications, whose general form is:
+{\tt $[$ ($root$) $]$ $string$ $[$ => $|$ <= $|$ == $]$ $[$ ($root$) $]$
+$string$}.
+
+This specifies a \rmindex{parse rule} ({\tt =>}) a \rmindex{print rule} ({\tt
+<=}) or both ({\tt ==}). The two $string$s preceded by optional parenthesized
+$root$s denote the left-hand and right-hand side of the rule as 'source
+code'.
+
+Rules are internalized wrt.\ an intermediate signature that is obtained from
+the parent theories' ones by adding all material of all sections preceding
+{\tt translations} in the {\tt .thy} file, especially new syntax defined in
+{\tt consts} is already effective.
+
+Then part of the process that transforms input strings into terms is applied:
+lexing, parsing and parse ast translations (see \S\ref{sec:asts}). Macros as
+specified in the parents are {\em not} expanded. Also note that the lexer
+runs in a different mode that additionally accepts identifiers of the form
+$\_~letter~quasiletter^*$ (like {\tt _idt}, {\tt _K}). The syntactic category
+to parse is specified by $root$, which defaults to {\tt logic}.
+
+Finally, Isabelle tries to guess which atoms of the resulting ast of the rule
+should be treated as constants during matching (see below). These names are
+extracted from all class, type and constant declarations made so far.
+
+\medskip
+The result are two lists of translation rules in internal form, that is pairs
+of asts. They can be viewed using {\tt Syntax.print_syntax}
+(\ttindex{parse_rules} and \ttindex{print_rules}). For {\tt SET} of
+Example~\ref{ex:set_trans} these are:
+\begin{ttbox}
+parse_rules:
+ ("{\at}Collect" x A P) -> ("Collect" A ("_abs" x P))
+ ("{\at}Replace" y x A Q) -> ("Replace" A ("_abs" x ("_abs" y Q)))
+ ("{\at}Ball" x A P) -> ("Ball" A ("_abs" x P))
+print_rules:
+ ("Collect" A ("_abs" x P)) -> ("{\at}Collect" x A P)
+ ("Replace" A ("_abs" x ("_abs" y Q))) -> ("{\at}Replace" y x A Q)
+ ("Ball" A ("_abs" x P)) -> ("{\at}Ball" x A P)
+\end{ttbox}
+
+Note that in this notation all rules are oriented left to right. In the {\tt
+translations} section, which has been split into two parts, print rules
+appeared right to left.
+
+\begin{warn}
+Be careful not to choose names for variables in rules that are actually
+treated as constant. If in doubt, check the rules in their internal form or
+the section labeled {\tt consts} in the output of {\tt Syntax.print_syntax}.
+\end{warn}
+
+
+\subsection{Applying rules}
+
+In the course of parsing and printing terms, asts are generated as an
+intermediate form as pictured in Figure~\ref{fig:parse_print}. These asts are
+normalized wrt.\ the given lists of translation rules in a uniform manner. As
+stated earlier, asts are supposed to be sort of first order terms. The
+rewriting systems derived from {\tt translations} sections essentially
+resemble traditional first order term rewriting systems. We'll first examine
+how a single rule is applied.
+
+Let $t$ be the ast to be normalized and $(l, r)$ some translation rule. A
+(improper) subast $u$ of $t$ is called {\bf redex}\indexbold{redex (ast)}
+(reducible expression), if it is an instance of $l$. In this case $l$ is said
+to {\bf match}\indexbold{match (ast)} $u$. A redex matched by $l$ may be
+replaced by the corresponding instance of $r$, thus {\bf
+rewriting}\index{rewrite (ast)} the ast $t$.
+
+Matching requires some notion of {\bf place-holders}\indexbold{place-holder
+(ast)} that may occur in rule patterns but not in ordinary asts, which are
+considered ground. Here we simply use {\tt Variable}s for this purpose.
+
+More formally, the matching of $u$ by $l$ is performed as follows (the rule
+pattern is the second argument): \index{match (ast)@$match$ (ast)}
+\begin{itemize}
+ \item $match(\Constant x, \Constant x) = \mbox{OK}$.
+
+ \item $match(\Variable x, \Constant x) = \mbox{OK}$.
+
+ \item $match(u, \Variable x) = \mbox{OK, bind}~x~\mbox{to}~u$.
+
+ \item $match(\Appl{u@1, \ldots, u@n}, \Appl{l@1, \ldots, l@n}) = match(u@1,
+ l@1), \ldots, match(u@n, l@n)$.
+
+ \item $match(u, l) = \mbox{FAIL}$ in any other case.
+\end{itemize}
+
+This means that a {\tt Constant} pattern matches any atomic asts of the same
+name, while a {\tt Variable} matches any ast. If successful, $match$ yields a
+substitution $\sigma$ for all place-holders in $l$, equalling it with the
+redex $u$. Then $\sigma$ is applied to $r$ generating the appropriate
+instance that replaces $u$.
+
+\medskip
+In order to make things simple and fast, ast rewrite rules $(l, r)$ are
+restricted by the following conditions:
+\begin{itemize}
+ \item Rules have to be left linear, i.e.\ $l$ must not contain any {\tt
+ Variable} more than once.
+
+ \item Rules must have constant heads, i.e.\ $l = \mtt"c\mtt"$ or $l =
+ (\mtt"c\mtt" ~ x@1 \ldots x@n)$.
+
+ \item The set of variables contained in $r$ has to be a subset of those of
+ $l$.
+\end{itemize}
+
+\medskip
+Having first order matching in mind, the second case of $match$ may look a
+bit odd. But this is exactly the place, where {\tt Variable}s of non-rule
+asts behave like {\tt Constant}s. The deeper meaning of this is related with
+asts being very 'primitive' in some sense, ignorant of the underlying
+'semantics', only short way from parse trees. At this level it is not yet
+known, which $id$s will become constants, bounds, frees, types or classes. As
+$ast_of_pt$ (see \S\ref{sec:asts}) shows, former parse tree heads appear in
+asts as {\tt Constant}s, while $id$s, $var$s, $tfree$s and $tvar$s become
+{\tt Variable}s.
+
+This is at variance with asts generated from terms before printing (see
+$ast_of_term$ in \S\ref{sec:asts}), where all constants and type constructors
+become {\tt Constant}s.
+
+\begin{warn}
+This means asts may contain quite a messy mixture of {\tt Variable}s and {\tt
+Constant}s, which is insignificant at macro level because $match$ treats them
+alike.
+\end{warn}
+
+Because of this behaviour, different kinds of atoms with the same name are
+indistinguishable, which may make some rules prone to misbehaviour. Regard
+the following fragmentary example:
+\begin{ttbox}
+types
+ Nil 0
+consts
+ Nil :: "'a list"
+ "[]" :: "'a list" ("[]")
+translations
+ "[]" == "Nil"
+\end{ttbox}
+Then the term {\tt Nil} will be printed as {\tt []}, just as expected. What
+happens with \verb|%Nil.t| or {\tt x::Nil} is left as an exercise.
+
+
+\subsection{Rewriting strategy}
+
+When normalizing an ast by repeatedly applying translation rules until no
+more rule is applicable, there are in each step two choices: which rule to
+apply next, and which redex to reduce.
+
+We could assume that the user always supplies terminating and confluent
+rewriting systems, but this would often complicate things unnecessarily.
+Therefore, we reveal part of the actual rewriting strategy: The normalizer
+always applies the first matching rule reducing an unspecified redex chosen
+first.
+
+Thereby, 'first rule' is roughly speaking meant wrt.\ the appearance of the
+rules in the {\tt translations} sections. But this is more tricky than it
+seems: If a given theory is {\em extended}, new rules are simply appended to
+the end. But if theories are {\em merged}, it is not clear which list of
+rules has priority over the other. In fact the merge order is left
+unspecified. This shouldn't cause any problems in practice, since
+translations of different theories usually do not overlap. {\tt
+Syntax.print_syntax} shows the rules in their internal order.
+
+\medskip
+You can watch the normalization of asts during parsing and printing by
+issuing: \index{*Syntax.trace_norm_ast}
+\begin{ttbox}
+Syntax.trace_norm_ast := true;
+\end{ttbox}
+An alternative is the use of \ttindex{Syntax.test_read}, which is always in
+trace mode. The information displayed when tracing\index{tracing (ast)}
+includes: the ast before normalization ({\tt pre}), redexes with results
+({\tt rewrote}), the normal form finally reached ({\tt post}) and some
+statistics ({\tt normalize}). If tracing is off,
+\ttindex{Syntax.stat_norm_ast} can be set to {\tt true} in order to enable
+printing of the normal form and statistics only.
+
+
+\subsection{More examples}
+
+Let us first reconsider Example~\ref{ex:set_trans}, which is concerned with
+variable binding constructs.
+
+There is a minor disadvantage over an implementation via translation
+functions (as done for binders):
+
+\begin{warn}
+If \ttindex{eta_contract} is set to {\tt true}, terms will be
+$\eta$-contracted {\em before} the ast rewriter sees them. Thus some
+abstraction nodes needed for print rules to match may get lost. E.g.\
+\verb|Ball(A, %x. P(x))| is contracted to {\tt Ball(A, P)}, the print rule is
+no longer applicable and the output will be {\tt Ball(A, P)}. Note that
+$\eta$-expansion via macros is {\em not} possible.
+\end{warn}
+
+\medskip
+Another common trap are meta constraints. If \ttindex{show_types} is set to
+{\tt true}, bound variables will be decorated by their meta types at the
+binding place (but not at occurrences in the body). E.g.\ matching with
+\verb|Collect(A, %x. P)| binds {\tt x} to something like {\tt ("_constrain" y
+"i")} rather than only {\tt y}. Ast rewriting will cause the constraint to
+appear in the external form, say \verb|{y::i:A::i. P::o}|. Therefore your
+syntax should be ready for such constraints to be re-read. This is the case
+in our example, because of the category {\tt idt} of the first argument.
+
+\begin{warn}
+Choosing {\tt id} instead of {\tt idt} is a very common error, especially
+since it appears in former versions of most of Isabelle's object-logics.
+\end{warn}
+
+\begin{example} \label{ex:finset_trans}
+This example demonstrates the use of recursive macros to implement a
+convenient notation for finite sets.
+\begin{ttbox}
+FINSET = SET +
+types
+ is 0
+consts
+ "" :: "i => is" ("_")
+ "{\at}Enum" :: "[i, is] => is" ("_,/ _")
+ empty :: "i" ("{\ttlbrace}{\ttrbrace}")
+ insert :: "[i, i] => i"
+ "{\at}Finset" :: "is => i" ("{\ttlbrace}(_){\ttrbrace}")
+translations
+ "{\ttlbrace}x, xs{\ttrbrace}" == "insert(x, {\ttlbrace}xs{\ttrbrace})"
+ "{\ttlbrace}x{\ttrbrace}" == "insert(x, {\ttlbrace}{\ttrbrace})"
+end
+\end{ttbox}
+
+Finite sets are internally built up by {\tt empty} and {\tt insert}.
+Externally we would like to see \verb|{x, y, z}| rather than {\tt insert(x,
+insert(y, insert(z, empty)))}.
+
+First we define the generic syntactic category {\tt is} for one or more
+objects of type {\tt i} separated by commas (including breaks for pretty
+printing). The category has to be declared as a 0-place type constructor, but
+without {\tt arities} declaration. Hence {\tt is} is not a logical type, no
+default productions will be added, and we can cook our own syntax for {\tt
+is} (first two lines of {\tt consts} section). If we had needed generic
+enumerations of type $\alpha$ (i.e.\ {\tt logic}), we could have used the
+predefined category \ttindex{args} and skipped this part altogether.
+
+Next follows {\tt empty}, which is already equipped with its syntax
+\verb|{}|, and {\tt insert} without concrete syntax. The syntactic constant
+{\tt\at Finset} provides concrete syntax for enumerations of {\tt i} enclosed
+in curly braces. Remember that a pair of parentheses specifies a block of
+indentation for pretty printing. The category {\tt is} can be later reused
+for other enumerations like lists or tuples.
+
+The translations might look a bit odd at the first glance. But rules can be
+only fully understood in their internal forms, which are:
+\begin{ttbox}
+parse_rules:
+ ("{\at}Finset" ("{\at}Enum" x xs)) -> ("insert" x ("{\at}Finset" xs))
+ ("{\at}Finset" x) -> ("insert" x "empty")
+print_rules:
+ ("insert" x ("{\at}Finset" xs)) -> ("{\at}Finset" ("{\at}Enum" x xs))
+ ("insert" x "empty") -> ("{\at}Finset" x)
+\end{ttbox}
+This shows that \verb|{x, xs}| indeed matches any set enumeration of at least
+two elements, binding the first to {\tt x} and the rest to {\tt xs}.
+Likewise, \verb|{xs}| and \verb|{x}| represent any set enumeration. Note that
+the parse rules only work in this order.
+
+\medskip
+Some rules are prone to misbehaviour, as
+\verb|%empty insert. insert(x, empty)| shows, which is printed as
+\verb|%empty insert. {x}|. This problem arises, because the ast rewriter
+cannot discern constants, frees, bounds etc.\ and looks only for names of
+atoms.
+
+Thus the names of {\tt Constant}s occurring in the (internal) left-hand side
+of translation rules should be regarded as 'reserved keywords'. It is good
+practice to choose non-identifiers here like {\tt\at Finset} or sufficiently
+long and strange names.
+\end{example}
+
+\begin{example} \label{ex:prod_trans}
+One of the well-formedness conditions for ast rewrite rules stated earlier
+implies that you can never introduce new {\tt Variable}s on the right-hand
+side. Something like \verb|"K(B)" => "%x. B"| is illegal and could cause
+variable capturing, if it were allowed. In such cases you usually have to
+fall back on translation functions. But there is a trick that makes things
+quite readable in some cases: {\em calling print translations by print
+rules}. This is demonstrated here.
+\begin{ttbox}
+PROD = FINSET +
+consts
+ Pi :: "[i, i => i] => i"
+ "{\at}PROD" :: "[idt, i, i] => i" ("(3PROD _:_./ _)" 10)
+ "{\at}->" :: "[i, i] => i" ("(_ ->/ _)" [51, 50] 50)
+translations
+ "PROD x:A. B" => "Pi(A, %x. B)"
+ "A -> B" => "Pi(A, _K(B))"
+end
+ML
+ val print_translation = [("Pi", dependent_tr' ("{\at}PROD", "{\at}->"))];
+\end{ttbox}
+
+{\tt Pi} is an internal constant for constructing dependent products. Two
+external forms exist: {\tt PROD x:A.B}, the general case, and {\tt A -> B},
+an abbreviation for \verb|Pi(A, %x.B)| with {\tt B} not actually depending on
+{\tt x}.
+
+Now the second parse rule is where the trick comes in: {\tt _K(B)} is
+introduced during ast rewriting, which later becomes \verb|%x.B| due to a
+parse translation associated with \ttindex{_K}. Note that a leading {\tt _}
+in $id$s is allowed in translation rules, but not in ordinary terms. This
+special behaviour of the lexer is very useful for 'forging' asts containing
+names that are not directly accessible normally. Unfortunately, there is no
+such trick for printing, so we have to add a {\tt ML} section for the print
+translation \ttindex{dependent_tr'}.
+
+The parse translation for {\tt _K} is already installed in Pure, and {\tt
+dependent_tr'} is exported by the syntax module for public use. See
+\S\ref{sec:tr_funs} for more of the arcane lore of translation functions.
+\end{example}
+
+
+
+\section{Translation functions *} \label{sec:tr_funs}
+
+This section is about the remaining translation mechanism which enables the
+designer of theories to do almost everything with terms or asts during the
+parsing or printing process, by writing some \ML-functions. The logic \LK\ is
+a good example of a quite sophisticated use of this to transform between
+internal and external representations of associative sequences. The high
+level macro system described in \S\ref{sec:macros} fails here completely.
+
+\begin{warn}
+A full understanding of the matters presented here requires some familiarity
+with Isabelle's internals, especially the datatypes {\tt term}, {\tt typ},
+{\tt Syntax.ast} and the encodings of types and terms as such at the various
+stages of the parsing or printing process. You probably do not really want to
+use translation functions at all!
+\end{warn}
+
+As already mentioned in \S\ref{sec:asts}, there are four kinds of translation
+functions. All such functions are associated with a name which specifies an
+ast's or term's head invoking that function. Such names can be (logical or
+syntactic) constants or type constructors.
+
+{\tt Syntax.print_syntax} displays the sets of names associated with the
+translation functions of a {\tt Syntax.syntax} under
+\ttindex{parse_ast_translation}, \ttindex{parse_translation},
+\ttindex{print_translation} and \ttindex{print_ast_translation}. The user can
+add new ones via the {\tt ML} section\index{ML section@{\tt ML} section} of a
+{\tt .thy} file. But there may never be more than one function of the same
+kind per name.
+
+\begin{warn}
+Conceptually, the {\tt ML} section should appear between {\tt consts} and
+{\tt translations}, i.e.\ newly installed translation functions are already
+effective when macros and logical rules are parsed. {\tt ML} has to be the
+last section because the {\tt .thy} file parser is unable to detect the end
+of \ML\ code in another way than by end-of-file.
+\end{warn}
+
+All text of the {\tt ML} section is simply copied verbatim into the \ML\ file
+generated from a {\tt .thy} file. Definitions made here by the user become
+components of a \ML\ structure of the same name as the theory to be created.
+Therefore local things should be declared within {\tt local}. The following
+special \ML\ values, which are all optional, serve as the interface for the
+installation of user defined translation functions.
+
+\begin{ttbox}
+val parse_ast_translation: (string * (ast list -> ast)) list
+val parse_translation: (string * (term list -> term)) list
+val print_translation: (string * (term list -> term)) list
+val print_ast_translation: (string * (ast list -> ast)) list
+\end{ttbox}
+
+The basic idea behind all four kinds of functions is relatively simple (see
+also Figure~\ref{fig:parse_print}): Whenever --- during the transformations
+between parse trees, asts and terms --- a combination of the form
+$(\mtt"c\mtt"~x@1 \ldots x@n)$ is encountered, and a translation function $f$
+of appropriate kind exists for $c$, the result will be $f \mtt[ x@1, \ldots,
+x@n \mtt]$. Thereby, $x@1, \ldots, x@n$ (with $n \ge 0$) are asts for ast
+translations and terms for term translations. A 'combination' at ast level is
+of the form $\Constant c$ or $\Appl{\Constant c, x@1, \ldots, x@n}$, and at
+term level $\ttfct{Const} (c, \tau)$ or $\ttfct{Const} (c, \tau) \ttrel{\$}
+x@1 \ttrel{\$} \dots \ttrel{\$} x@n$.
+
+\medskip
+Translation functions at ast level differ from those at term level only in
+the same way, as asts and terms differ. Terms, being more complex and more
+specific, allow more sophisticated transformations (typically involving
+abstractions and bound variables).
+
+On the other hand, {\em parse} (ast) translations differ from {\em print}
+(ast) translations more fundamentally:
+\begin{description}
+ \item[Parse (ast) translations] are applied bottom-up, i.e.\ the arguments
+ supplied ($x@1, \ldots, x@n$ above) are already in translated form.
+ Additionally, they may not fail, exceptions are re-raised after printing
+ of an error message.
+
+ \item[Print (ast) translations] are applied top-down, i.e.\ supplied with
+ arguments that are partly still in internal form. The result is again fed
+ into the translation machinery as a whole. Therefore a print (ast)
+ translation should not introduce as head a constant of the same name that
+ invoked it in the first place. Alternatively, exception \ttindex{Match}
+ may be raised, indicating failure of translation.
+\end{description}
+
+Another difference between the parsing and the printing process is, which
+atoms are {\tt Constant}s or {\tt Const}s, i.e.\ able to invoke translation
+functions.
+
+For parse ast translations only former parse tree heads are {\tt Constant}s
+(see also $ast_of_pt$ in \S\ref{sec:asts}). These and additionally introduced
+{\tt Constant}s (e.g.\ by macros), become {\tt Const}s for parse translations
+(see also $term_of_ast$ in \S\ref{sec:asts}).
+
+The situation is slightly different, when terms are prepared for printing,
+since the role of atoms is known. Initially, all logical constants and type
+constructors may invoke print translations. New constants may be introduced
+by these or by macros, able to invoke parse ast translations.
+
+
+\subsection{A simple example *}