Initial revision
authorlcp
Wed, 10 Nov 1993 05:00:57 +0100
changeset 104 d8205bb279a7
parent 103 30bd42401ab2
child 105 216d6ed87399
Initial revision
doc-src/Logics/CTT-eg.txt
doc-src/Logics/CTT-rules.txt
doc-src/Logics/CTT.tex
doc-src/Logics/FOL-eg.txt
doc-src/Logics/FOL.tex
doc-src/Logics/HOL-eg.txt
doc-src/Logics/HOL-rules.txt
doc-src/Logics/LK-eg.txt
doc-src/Logics/LK.tex
doc-src/Logics/Old_HOL.tex
doc-src/Logics/ZF-eg.txt
doc-src/Logics/ZF-rules.txt
doc-src/Logics/ZF.tex
doc-src/Logics/abstract.txt
doc-src/Logics/defining.tex
doc-src/Logics/intro.tex
doc-src/Logics/logics.bbl
doc-src/Logics/logics.tex
doc-src/Logics/logics.toc
doc-src/Logics/old.defining.tex
doc-src/Ref/abstract.txt
doc-src/Ref/classical.tex
doc-src/Ref/goals.tex
doc-src/Ref/introduction.tex
doc-src/Ref/ref.bbl
doc-src/Ref/ref.tex
doc-src/Ref/ref.toc
doc-src/Ref/simp.tex
doc-src/Ref/simplifier-eg.txt
doc-src/Ref/simplifier.tex
doc-src/Ref/substitution.tex
doc-src/Ref/tactic.tex
doc-src/Ref/tctical.tex
doc-src/Ref/theories.tex
doc-src/Ref/thm.tex
doc-src/Ref/undocumented.tex
doc-src/ind-defs.bbl
doc-src/ind-defs.toc
--- /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 *}
+
+Presenting a simple and useful example of translation functions is not that
+easy, since the macro system is sufficient for most simple applications. By
+convention, translation functions always have names ending with {\tt
+_ast_tr}, {\tt _tr}, {\tt _tr'} or {\tt _ast_tr'}. You may look for such
+names in the sources of Pure Isabelle for more examples.
+
+\begin{example} \label{ex:tr_funs}
+
+We continue Example~\ref{ex:prod_trans} by presenting the \ML\ sources of the
+parse translation for \ttindex{_K} and the print translation
+\ttindex{dependent_tr'}:
+
+\begin{ttbox}
+(* nondependent abstraction *)
+
+fun k_tr (*"_K"*) [t] = Abs ("x", dummyT, incr_boundvars 1 t)
+  | k_tr (*"_K"*) ts = raise_term "k_tr" ts;
+
+(* dependent / nondependent quantifiers *)
+
+fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
+      if 0 mem (loose_bnos B) then
+        let val (x', B') = variant_abs (x, dummyT, B);
+        in list_comb (Const (q, dummyT) $ Free (x', T) $ A $ B', ts)
+        end
+      else list_comb (Const (r, dummyT) $ A $ B, ts)
+  | dependent_tr' _ _ = raise Match;
+\end{ttbox}
+
+This text is taken from the Pure sources, ordinary user translations would
+typically appear within the {\tt ML} section of the {\tt .thy} file.
+
+\medskip
+If {\tt k_tr} is called with exactly one argument $t$, it creates a new {\tt
+Abs} node with a body derived from $t$: loose {\tt Bound}s, i.e.\ those
+referring to outer {\tt Abs}s, are incremented using
+\ttindex{incr_boundvars}. This and many other basic term manipulation
+functions are defined in {\tt Pure/term.ML}, the comments there being in most
+cases the only documentation.
+
+Since terms fed into parse translations are not yet typed, the type of the
+bound variable in the new {\tt Abs} is simply {\tt dummyT}.
+
+\medskip
+The argument $(q, r)$ for {\tt dependent_tr'} is supplied already during the
+installation within a {\tt ML} section. This yields a parse translation that
+transforms something like $c(A, \mtt{Abs}(x, T, B), t@1, \ldots, t@n)$ into
+$q(x', A, B', t@1, \ldots, t@n)$ or $r(A, B, t@1, \ldots, t@n)$. The latter
+form, if $B$ does not actually depend on $x$. This is checked using
+\ttindex{loose_bnos}, yet another undocumented function of {\tt
+Pure/term.ML}. Note that $x'$ is a version of $x$ renamed away from all names
+in $B$, and $B'$ the body $B$ with {\tt Bound}s referring to our {\tt Abs}
+node replaced by $\ttfct{Free} (x', \mtt{dummyT})$.
+
+We have to be more careful with types here. While types of {\tt Const}s are
+completely ignored, type constraints may be printed for some {\tt Free}s and
+{\tt Var}s (only if \ttindex{show_types} is set to {\tt true}). Variables of
+type \ttindex{dummyT} are never printed with constraint, though. Thus, a
+constraint of $x'$ may only appear at its binding place, since {\tt Free}s of
+$B'$ replacing the appropriate {\tt Bound}s of $B$ via \ttindex{variant_abs}
+have all type {\tt dummyT}.
+\end{example}
+
+
+
+\section{Example: some minimal logics} \label{sec:min_logics}
+
+This concluding section presents some examples that are very simple from a
+syntactic point of view. They should rather demonstrate how to define new
+object-logics from scratch. In particular we need to say how an object-logic
+syntax is hooked up to the meta-logic. Since all theorems must conform to the
+syntax for $prop$ (see Figure~\ref{fig:pure_gram}), that syntax has to be
+extended with the object-level syntax. Assume that the syntax of your
+object-logic defines a category $o$ of formulae. These formulae can now
+appear in axioms and theorems wherever $prop$ does if you add the production
+\[ prop ~=~ o. \]
+More precisely, you need a coercion from formulae to propositions:
+\begin{ttbox}
+Base = Pure +
+types
+  o 0
+arities
+  o :: logic
+consts
+  Trueprop :: "o => prop"   ("_" 5)
+end
+\end{ttbox}
+The constant {\tt Trueprop} (the name is arbitrary) acts as an invisible
+coercion function. Assuming this definition resides in a file {\tt base.thy},
+you have to load it with the command {\tt use_thy "base"}.
+
+One of the simplest nontrivial logics is {\em minimal logic} of implication.
+Its definition in Isabelle needs no advanced features but illustrates the
+overall mechanism quite nicely:
+\begin{ttbox}
+Hilbert = Base +
+consts
+  "-->" :: "[o, o] => o"   (infixr 10)
+rules
+  K     "P --> Q --> P"
+  S     "(P --> Q --> R) --> (P --> Q) --> P --> R"
+  MP    "[| P --> Q; P |] ==> Q"
+end
+\end{ttbox}
+After loading this definition you can start to prove theorems in this logic:
+\begin{ttbox}
+goal Hilbert.thy "P --> P";
+{\out Level 0}
+{\out P --> P}
+{\out  1.  P --> P}
+by (resolve_tac [Hilbert.MP] 1);
+{\out Level 1}
+{\out P --> P}
+{\out  1.  ?P --> P --> P}
+{\out  2.  ?P}
+by (resolve_tac [Hilbert.MP] 1);
+{\out Level 2}
+{\out P --> P}
+{\out  1.  ?P1 --> ?P --> P --> P}
+{\out  2.  ?P1}
+{\out  3.  ?P}
+by (resolve_tac [Hilbert.S] 1);
+{\out Level 3}
+{\out P --> P}
+{\out  1.  P --> ?Q2 --> P}
+{\out  2.  P --> ?Q2}
+by (resolve_tac [Hilbert.K] 1);
+{\out Level 4}
+{\out P --> P}
+{\out  1.  P --> ?Q2}
+by (resolve_tac [Hilbert.K] 1);
+{\out Level 5}
+{\out P --> P}
+{\out No subgoals!}
+\end{ttbox}
+As you can see, this Hilbert-style formulation of minimal logic is easy to
+define but difficult to use. The following natural deduction formulation is
+far preferable:
+\begin{ttbox}
+MinI = Base +
+consts
+  "-->" :: "[o, o] => o"   (infixr 10)
+rules
+  impI  "(P ==> Q) ==> P --> Q"
+  impE  "[| P --> Q; P |] ==> Q"
+end
+\end{ttbox}
+Note, however, that although the two systems are equivalent, this fact cannot
+be proved within Isabelle: {\tt S} and {\tt K} can be derived in {\tt MinI}
+(exercise!), but {\tt impI} cannot be derived in {\tt Hilbert}. The reason is
+that {\tt impI} is only an {\em admissible} rule in {\tt Hilbert}, something
+that can only be shown by induction over all possible proofs in {\tt
+Hilbert}.
+
+It is a very simple matter to extend minimal logic with falsity:
+\begin{ttbox}
+MinIF = MinI +
+consts
+  False :: "o"
+rules
+  FalseE "False ==> P"
+end
+\end{ttbox}
+On the other hand, we may wish to introduce conjunction only:
+\begin{ttbox}
+MinC = Base +
+consts
+  "&" :: "[o, o] => o"   (infixr 30)
+rules
+  conjI  "[| P; Q |] ==> P & Q"
+  conjE1 "P & Q ==> P"
+  conjE2 "P & Q ==> Q"
+end
+\end{ttbox}
+And if we want to have all three connectives together, we define:
+\begin{ttbox}
+MinIFC = MinIF + MinC
+\end{ttbox}
+Now we can prove mixed theorems like
+\begin{ttbox}
+goal MinIFC.thy "P & False --> Q";
+by (resolve_tac [MinI.impI] 1);
+by (dresolve_tac [MinC.conjE2] 1);
+by (eresolve_tac [MinIF.FalseE] 1);
+\end{ttbox}
+Try this as an exercise!
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/intro.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,139 @@
+%% $Id$
+\chapter{Introduction}
+Several logics come with Isabelle.  Many of them are sufficiently developed
+to serve as comfortable reasoning environments.  They are also good
+starting points for defining new logics.  Each logic is distributed with
+sample proofs, some of which are described in this document.
+
+\begin{quote}
+{\ttindexbold{FOL}} is many-sorted first-order logic with natural
+deduction.  It comes in both constructive and classical versions.
+
+{\ttindexbold{ZF}} is axiomatic set theory, using the Zermelo-Fraenkel
+axioms~\cite{suppes72}.  It is built upon classical~\FOL{}.
+ 
+{\ttindexbold{HOL}} is the higher-order logic of Church~\cite{church40},
+which is also implemented by Gordon's~{\sc hol} system~\cite{gordon88a}.  This
+object-logic should not be confused with Isabelle's meta-logic, which is
+also a form of higher-order logic.
+ 
+{\ttindexbold{CTT}} is a version of Martin-L\"of's Constructive Type
+Theory~\cite{nordstrom90}, with extensional equality.  Universes are not
+included.
+ 
+{\ttindexbold{LK}} is another version of first-order logic, a classical
+sequent calculus.  Sequents have the form $A@1,\ldots,A@m\turn
+B@1,\ldots,B@n$; rules are applied using associative matching.
+
+{\ttindexbold{Modal}} implements the modal logics $T$, $S4$, and~$S43$.  It
+is built upon~\LK{}.
+
+{\ttindexbold{Cube}} is Barendregt's $\lambda$-cube.
+
+{\ttindexbold{LCF}} is a version of Scott's Logic for Computable Functions,
+which is also implemented by the~{\sc lcf} system~\cite{paulson87}.
+\end{quote}
+The logics {\tt Modal}, {\tt Cube} and {\tt LCF} are currently undocumented.
+
+This manual assumes that you have read the {\em Introduction to Isabelle\/}
+and have some experience of using Isabelle to perform interactive proofs.
+It refers to packages defined in the {\em Reference Manual}, which you
+should keep at hand.
+
+
+\section{Syntax definitions}
+This manual defines each logic's syntax using a context-free grammar.
+These grammars obey the following conventions:
+\begin{itemize}
+\item identifiers denote nonterminal symbols
+\item {\tt typewriter} font denotes terminal symbols
+\item parentheses $(\ldots)$ express grouping
+\item constructs followed by a Kleene star, such as $id^*$ and $(\ldots)^*$
+can be repeated~0 or more times 
+\item alternatives are separated by a vertical bar,~$|$
+\item the symbol for alphanumeric identifier is~{\it id\/} 
+\item the symbol for scheme variables is~{\it var}
+\end{itemize}
+To reduce the number of nonterminals and grammar rules required, Isabelle's
+syntax module employs {\bf precedences}.  Each grammar rule is given by a
+mixfix declaration, which has a precedence, and each argument place has a
+precedence.  This general approach handles infix operators that associate
+either to the left or to the right, as well as prefix and binding
+operators.
+
+In a syntactically valid expression, an operator's arguments never involve
+an operator of lower precedence unless brackets are used.  Consider
+first-order logic, where $\exists$ has lower precedence than $\disj$,
+which has lower precedence than $\conj$.  There, $P\conj Q \disj R$
+abbreviates $(P\conj Q) \disj R$ rather than $P\conj (Q\disj R)$.  Also,
+$\exists x.P\disj Q$ abbreviates $\exists x.(P\disj Q)$ rather than
+$(\exists x.P)\disj Q$.  Note especially that $P\disj(\exists x.Q)$
+becomes syntactically invalid if the brackets are removed.
+
+A {\bf binder} is a symbol associated with a constant of type
+$(\sigma\To\tau)\To\tau'$.  For instance, we may declare~$\forall$ as a
+binder for the constant~$All$, which has type $(\alpha\To o)\To o$.  This
+defines the syntax $\forall x.t$ to mean $\forall(\lambda x.t)$.  We can
+also write $\forall x@1\ldots x@m.t$ to abbreviate $\forall x@1.  \ldots
+\forall x@m.t$; this is possible for any constant provided that $\tau$ and
+$\tau'$ are the same type.  \HOL's description operator $\epsilon x.P(x)$
+has type $(\alpha\To bool)\To\alpha$ and can bind only one variable, except
+when $\alpha$ is $bool$.  \ZF's bounded quantifier $\forall x\in A.P(x)$
+cannot be declared as a binder because it has type $[i, i\To o]\To o$.  The
+syntax for binders allows type constraints on bound variables, as in
+\[ \forall (x{::}\alpha) \; y{::}\beta. R(x,y) \]
+
+To avoid excess detail, the logic descriptions adopt a semi-formal style.
+Infix operators and binding operators are listed in separate tables, which
+include their precedences.  Grammars do not give numeric precedences;
+instead, the rules appear in order of decreasing precedence.  This should
+suffice for most purposes; for detailed precedence information, consult the
+syntax definitions in the {\tt.thy} files.  Chapter~\ref{Defining-Logics}
+describes Isabelle's syntax module, including its use of precedences.
+
+Each nonterminal symbol is associated with some Isabelle type.  For
+example, the {\bf formulae} of first-order logic have type~$o$.  Every
+Isabelle expression of type~$o$ is therefore a formula.  These include
+atomic formulae such as $P$, where $P$ is a variable of type~$o$, and more
+generally expressions such as $P(t,u)$, where $P$, $t$ and~$u$ have
+suitable types.  Therefore, ``expression of type~$o$'' is listed as a
+separate possibility in the grammar for formulae.
+
+Infix operators are represented internally by constants with the prefix
+\hbox{\tt"op "}.  For instance, implication is the constant
+\hbox{\tt"op~-->"}.  This permits infixes to be used in non-infix contexts
+(just as with \ML's~{\tt op} keyword).  You also need to know the name of
+the internal constant if you are writing code that inspects terms.
+
+
+\section{Proof procedures}
+Most object-logics come with simple proof procedures.  These are reasonably
+powerful for interactive use, though often simplistic and incomplete.  You
+can do single-step proofs using \verb|resolve_tac| and
+\verb|assume_tac|, referring to the inference rules of the logic by {\sc
+ml} identifiers.
+
+Call a rule {\em safe\/} if applying it to a provable goal always yields
+provable subgoals.  If a rule is safe then it can be applied automatically
+to a goal without destroying our chances of finding a proof.  For instance,
+all the rules of the classical sequent calculus {\sc lk} are safe.
+Intuitionistic logic includes some unsafe rules, like disjunction
+introduction ($P\disj Q$ can be true when $P$ is false) and existential
+introduction ($\ex{x}P(x)$ can be true when $P(a)$ is false for certain
+$a$).  Universal elimination is unsafe if the formula $\all{x}P(x)$ is
+deleted after use.
+
+Proof procedures use safe rules whenever possible, delaying the application
+of unsafe rules. Those safe rules are preferred that generate the fewest
+subgoals. Safe rules are (by definition) deterministic, while the unsafe
+rules require search. The design of a suitable set of rules can be as
+important as the strategy for applying them.
+
+Many of the proof procedures use backtracking.  Typically they attempt to
+solve subgoal~$i$ by repeatedly applying a certain tactic to it.  This
+tactic, which is known as a {\it step tactic}, resolves a selection of
+rules with subgoal~$i$. This may replace one subgoal by many; but the
+search persists until there are fewer subgoals in total than at the start.
+Backtracking happens when the search reaches a dead end: when the step
+tactic fails.  Alternative outcomes are then searched by a depth-first or
+best-first strategy.  
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/logics.bbl	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,140 @@
+\begin{thebibliography}{10}
+
+\bibitem{andrews86}
+Peter~B. Andrews.
+\newblock {\em An Introduction to Mathematical Logic and Type Theory: To Truth
+  Through Proof}.
+\newblock Academic Press, 1986.
+
+\bibitem{boyer86}
+Robert Boyer, Ewing Lusk, William McCune, Ross Overbeek, Mark Stickel, and
+  Lawrence Wos.
+\newblock Set theory in first-order logic: Clauses for {G\"odel's} axioms.
+\newblock {\em Journal of Automated Reasoning}, 2:287--327, 1986.
+
+\bibitem{church40}
+Alonzo Church.
+\newblock A formulation of the simple theory of types.
+\newblock {\em Journal of Symbolic Logic}, 5:56--68, 1940.
+
+\bibitem{dummett}
+Michael Dummett.
+\newblock {\em Elements of Intuitionism}.
+\newblock Oxford, 1977.
+
+\bibitem{dyckhoff}
+Roy Dyckhoff.
+\newblock Contraction-free sequent calculi for intuitionistic logic.
+\newblock {\em Journal of Symbolic Logic}, 57(3):795--807, 1992.
+
+\bibitem{felty91a}
+Amy Felty.
+\newblock A logic program for transforming sequent proofs to natural deduction
+  proofs.
+\newblock In Peter Schroeder-Heister, editor, {\em Extensions of Logic
+  Programming}, pages 157--178. Springer, 1991.
+\newblock LNAI 475.
+
+\bibitem{OBJ}
+K.~Futatsugi, J.A. Goguen, Jean-Pierre Jouannaud, and J.~Meseguer.
+\newblock Principles of {OBJ2}.
+\newblock In {\em Symposium on Principles of Programming Languages}, pages
+  52--66, 1985.
+
+\bibitem{gallier86}
+J.~H. Gallier.
+\newblock {\em Logic for Computer Science: Foundations of Automatic Theorem
+  Proving}.
+\newblock Harper \& Row, 1986.
+
+\bibitem{gordon88a}
+Michael J.~C. Gordon.
+\newblock {HOL}: A proof generating system for higher-order logic.
+\newblock In Graham Birtwistle and P.~A. Subrahmanyam, editors, {\em {VLSI}
+  Specification, Verification and Synthesis}, pages 73--128. Kluwer Academic
+  Publishers, 1988.
+
+\bibitem{halmos60}
+Paul~R. Halmos.
+\newblock {\em Naive Set Theory}.
+\newblock Van Nostrand, 1960.
+
+\bibitem{huet78}
+G.~P. Huet and B.~Lang.
+\newblock Proving and applying program transformations expressed with
+  second-order patterns.
+\newblock {\em Acta Informatica}, 11:31--55, 1978.
+
+\bibitem{martinlof84}
+Per Martin-L\"of.
+\newblock {\em Intuitionistic type theory}.
+\newblock Bibliopolis, 1984.
+
+\bibitem{nipkow-prehofer}
+Tobias Nipkow and Christian Prehofer.
+\newblock Type checking type classes.
+\newblock In {\em 20th ACM Symp.\ Principles of Programming Languages}, 1993.
+\newblock To appear.
+
+\bibitem{noel}
+Philippe {No\"el}.
+\newblock Experimenting with {Isabelle} in {ZF} set theory.
+\newblock {\em Journal of Automated Reasoning}.
+\newblock In press.
+
+\bibitem{nordstrom90}
+Bengt {Nordstr\"om}, Kent Petersson, and Jan Smith.
+\newblock {\em Programming in {Martin-L\"of}'s Type Theory. An Introduction}.
+\newblock Oxford, 1990.
+
+\bibitem{paulson87}
+Lawrence~C. Paulson.
+\newblock {\em Logic and Computation: Interactive proof with Cambridge LCF}.
+\newblock Cambridge University Press, 1987.
+
+\bibitem{paulson-COLOG}
+Lawrence~C. Paulson.
+\newblock A formulation of the simple theory of types (for {Isabelle}).
+\newblock In P.~Martin-L\"of and G.~Mints, editors, {\em COLOG-88:
+  International Conference on Computer Logic}, Tallinn, 1990. Estonian Academy
+  of Sciences, Springer.
+\newblock LNCS 417.
+
+\bibitem{pelletier86}
+F.~J. Pelletier.
+\newblock Seventy-five problems for testing automatic theorem provers.
+\newblock {\em Journal of Automated Reasoning}, 2:191--216, 1986.
+\newblock Errata, JAR 4 (1988), 235--236.
+
+\bibitem{plaisted90}
+David~A. Plaisted.
+\newblock A sequent-style model elimination strategy and a positive refinement.
+\newblock {\em Journal of Automated Reasoning}, 6(4):389--402, 1990.
+
+\bibitem{quaife92}
+Art Quaife.
+\newblock Automated deduction in {von Neumann-Bernays-G\"odel} set theory.
+\newblock {\em Journal of Automated Reasoning}, 8(1):91--147, 1992.
+
+\bibitem{suppes72}
+Patrick Suppes.
+\newblock {\em Axiomatic Set Theory}.
+\newblock Dover, 1972.
+
+\bibitem{takeuti87}
+G.~Takeuti.
+\newblock {\em Proof Theory}.
+\newblock North Holland, 2nd edition, 1987.
+
+\bibitem{thompson91}
+Simon Thompson.
+\newblock {\em Type Theory and Functional Programming}.
+\newblock Addison-Wesley, 1991.
+
+\bibitem{principia}
+A.~N. Whitehead and B.~Russell.
+\newblock {\em Principia Mathematica}.
+\newblock Cambridge University Press, 1962.
+\newblock Paperback edition to *56, abridged from the 2nd edition (1927).
+
+\end{thebibliography}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/logics.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,61 @@
+\documentstyle[a4,12pt,proof,iman,alltt]{report}
+%% $Id$
+%%%STILL NEEDS MODAL, LCF
+\includeonly{HOL}
+%%% to index derived rls:  ^\([a-zA-Z0-9][a-zA-Z0-9_]*\)        \\idx{\1}  
+%%% to index rulenames:   ^ *(\([a-zA-Z0-9][a-zA-Z0-9_]*\),     \\idx{\1}  
+%%% to index constants:   \\tt \([a-zA-Z0-9][a-zA-Z0-9_]*\)     \\idx{\1}  
+%%% to deverbify:         \\verb|\([^|]*\)|     \\ttindex{\1}  
+%% run    ../sedindex logics    to prepare index file
+\title{Isabelle's Object-Logics}
+
+\author{{\em Lawrence C. Paulson}\thanks{Tobias Nipkow and Markus Wenzel,
+    of T. U. Munich, wrote the chapter `Defining Logics'.  Markus Wenzel
+    also suggested changes and corrections.  Philippe de Groote wrote the
+    first version of the logic~\LK{} and contributed to~\ZF{}.  Tobias
+    Nipkow developed~\HOL{}, \LCF{} and~\Cube{}.  Philippe No\"el and
+    Martin Coen made many contributions to~\ZF{}.  Martin Coen
+    developed~\Modal{} with assistance from Rajeev Gor\'e.  The research
+    has been funded by the SERC (grants GR/G53279, GR/H40570) and by ESPRIT
+    project 6453: Types}.\\ 
+  Computer Laboratory \\ 
+  University of Cambridge \\[2ex] 
+  {\small{\em Electronic mail\/}: {\tt lcp@cl.cam.ac.uk}} \\[3cm]
+  {\small Copyright \copyright{} \number\year{} by Lawrence C. Paulson} }
+
+\date{}
+
+\newcommand\subcaption[1]{\par {\centering\normalsize\sc#1\par}\bigskip
+  \hrule\bigskip}
+
+\makeindex
+
+%For indexing names in ttbox; could be local to Chapters, making a subitem...
+\let\idx=\ttindexbold
+%%%%\newcommand\idx[1]{\indexbold{*#1}#1}
+
+\underscoreoff
+
+\setcounter{secnumdepth}{1} \setcounter{tocdepth}{2}  %% {secnumdepth}{2}???
+
+\pagestyle{headings}
+\sloppy
+\binperiod     %%%treat . like a binary operator
+
+\begin{document}
+\maketitle 
+\pagenumbering{roman} \tableofcontents \clearfirst
+\include{intro}
+\include{FOL}
+\include{ZF}
+\include{HOL}
+\include{LK}
+%%\include{Modal}
+\include{CTT}
+%%\include{Cube}
+%%\include{LCF}
+\include{defining}
+\bibliographystyle{plain}
+\bibliography{atp,theory,funprog,logicprog}
+\input{logics.ind}
+\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/logics.toc	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,100 @@
+\contentsline {chapter}{\numberline {1}Introduction}{1}
+\contentsline {section}{\numberline {1.1}Syntax definitions}{1}
+\contentsline {section}{\numberline {1.2}Proof procedures}{3}
+\contentsline {chapter}{\numberline {2}First-order logic}{4}
+\contentsline {section}{\numberline {2.1}Syntax and rules of inference}{4}
+\contentsline {section}{\numberline {2.2}Generic packages}{8}
+\contentsline {section}{\numberline {2.3}Intuitionistic proof procedures}{8}
+\contentsline {section}{\numberline {2.4}Classical proof procedures}{10}
+\contentsline {section}{\numberline {2.5}An intuitionistic example}{11}
+\contentsline {section}{\numberline {2.6}An example of intuitionistic negation}{12}
+\contentsline {section}{\numberline {2.7}A classical example}{14}
+\contentsline {section}{\numberline {2.8}Derived rules and the classical tactics}{16}
+\contentsline {subsection}{Deriving the introduction rule}{17}
+\contentsline {subsection}{Deriving the elimination rule}{17}
+\contentsline {subsection}{Using the derived rules}{18}
+\contentsline {subsection}{Derived rules versus definitions}{20}
+\contentsline {chapter}{\numberline {3}Zermelo-Fraenkel set theory}{23}
+\contentsline {section}{\numberline {3.1}Which version of axiomatic set theory?}{23}
+\contentsline {section}{\numberline {3.2}The syntax of set theory}{25}
+\contentsline {section}{\numberline {3.3}Binding operators}{25}
+\contentsline {section}{\numberline {3.4}The Zermelo-Fraenkel axioms}{28}
+\contentsline {section}{\numberline {3.5}From basic lemmas to function spaces}{33}
+\contentsline {subsection}{Fundamental lemmas}{33}
+\contentsline {subsection}{Unordered pairs and finite sets}{36}
+\contentsline {subsection}{Subset and lattice properties}{36}
+\contentsline {subsection}{Ordered pairs}{37}
+\contentsline {subsection}{Relations}{37}
+\contentsline {subsection}{Functions}{40}
+\contentsline {section}{\numberline {3.6}Further developments}{40}
+\contentsline {section}{\numberline {3.7}Simplification rules}{47}
+\contentsline {section}{\numberline {3.8}The examples directory}{48}
+\contentsline {section}{\numberline {3.9}A proof about powersets}{49}
+\contentsline {section}{\numberline {3.10}Monotonicity of the union operator}{51}
+\contentsline {section}{\numberline {3.11}Low-level reasoning about functions}{52}
+\contentsline {chapter}{\numberline {4}Higher-order logic}{55}
+\contentsline {section}{\numberline {4.1}Syntax}{55}
+\contentsline {subsection}{Types}{55}
+\contentsline {subsection}{Binders}{58}
+\contentsline {section}{\numberline {4.2}Rules of inference}{58}
+\contentsline {section}{\numberline {4.3}Generic packages}{62}
+\contentsline {section}{\numberline {4.4}A formulation of set theory}{63}
+\contentsline {subsection}{Syntax of set theory}{63}
+\contentsline {subsection}{Axioms and rules of set theory}{69}
+\contentsline {subsection}{Derived rules for sets}{69}
+\contentsline {section}{\numberline {4.5}Types}{69}
+\contentsline {subsection}{Product and sum types}{74}
+\contentsline {subsection}{The type of natural numbers, $nat$}{74}
+\contentsline {subsection}{The type constructor for lists, $\alpha \pcomma list$}{74}
+\contentsline {subsection}{The type constructor for lazy lists, $\alpha \pcomma llist$}{78}
+\contentsline {section}{\numberline {4.6}Classical proof procedures}{78}
+\contentsline {section}{\numberline {4.7}The examples directory}{78}
+\contentsline {section}{\numberline {4.8}Example: deriving the conjunction rules}{79}
+\contentsline {subsection}{The introduction rule}{79}
+\contentsline {subsection}{The elimination rule}{80}
+\contentsline {section}{\numberline {4.9}Example: Cantor's Theorem}{81}
+\contentsline {chapter}{\numberline {5}First-order sequent calculus}{83}
+\contentsline {section}{\numberline {5.1}Unification for lists}{83}
+\contentsline {section}{\numberline {5.2}Syntax and rules of inference}{84}
+\contentsline {section}{\numberline {5.3}Tactics for the cut rule}{84}
+\contentsline {section}{\numberline {5.4}Tactics for sequents}{88}
+\contentsline {section}{\numberline {5.5}Packaging sequent rules}{89}
+\contentsline {section}{\numberline {5.6}Proof procedures}{89}
+\contentsline {subsection}{Method A}{90}
+\contentsline {subsection}{Method B}{90}
+\contentsline {section}{\numberline {5.7}A simple example of classical reasoning}{91}
+\contentsline {section}{\numberline {5.8}A more complex proof}{92}
+\contentsline {chapter}{\numberline {6}Constructive Type Theory}{95}
+\contentsline {section}{\numberline {6.1}Syntax}{96}
+\contentsline {section}{\numberline {6.2}Rules of inference}{96}
+\contentsline {section}{\numberline {6.3}Rule lists}{101}
+\contentsline {section}{\numberline {6.4}Tactics for subgoal reordering}{104}
+\contentsline {section}{\numberline {6.5}Rewriting tactics}{105}
+\contentsline {section}{\numberline {6.6}Tactics for logical reasoning}{105}
+\contentsline {section}{\numberline {6.7}A theory of arithmetic}{106}
+\contentsline {section}{\numberline {6.8}The examples directory}{106}
+\contentsline {section}{\numberline {6.9}Example: type inference}{108}
+\contentsline {section}{\numberline {6.10}An example of logical reasoning}{109}
+\contentsline {section}{\numberline {6.11}Example: deriving a currying functional}{112}
+\contentsline {section}{\numberline {6.12}Example: proving the Axiom of Choice}{113}
+\contentsline {chapter}{\numberline {7}Defining Logics}{118}
+\contentsline {section}{\numberline {7.1}Precedence grammars}{118}
+\contentsline {section}{\numberline {7.2}Basic syntax}{119}
+\contentsline {subsection}{Logical types and default syntax}{120}
+\contentsline {subsection}{Lexical matters *}{121}
+\contentsline {subsection}{Inspecting syntax *}{121}
+\contentsline {section}{\numberline {7.3}Abstract syntax trees}{123}
+\contentsline {subsection}{Parse trees to asts}{125}
+\contentsline {subsection}{Asts to terms *}{126}
+\contentsline {subsection}{Printing of terms *}{126}
+\contentsline {section}{\numberline {7.4}Mixfix declarations}{128}
+\contentsline {subsection}{Infixes}{130}
+\contentsline {subsection}{Binders}{130}
+\contentsline {section}{\numberline {7.5}Syntactic translations (macros)}{131}
+\contentsline {subsection}{Specifying macros}{132}
+\contentsline {subsection}{Applying rules}{133}
+\contentsline {subsection}{Rewriting strategy}{135}
+\contentsline {subsection}{More examples}{135}
+\contentsline {section}{\numberline {7.6}Translation functions *}{138}
+\contentsline {subsection}{A simple example *}{139}
+\contentsline {section}{\numberline {7.7}Example: some minimal logics}{140}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Logics/old.defining.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,871 @@
+\chapter{Defining Logics} \label{Defining-Logics}
+This chapter is intended for Isabelle experts.  It explains how to define new
+logical systems, Isabelle's {\it raison d'\^etre}.  Isabelle logics are
+hierarchies of theories.  A number of simple examples are contained in the
+introductory manual; the full syntax of theory definitions is shown in the
+{\em Reference Manual}.  The purpose of this chapter is to explain the
+remaining subtleties, especially some context conditions on the class
+structure and the definition of new mixfix syntax.  A full understanding of
+the material requires knowledge of the internal representation of terms (data
+type {\tt term}) as detailed in the {\em Reference Manual}.  Sections marked
+with a * can be skipped on first reading.
+
+
+\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 :: (\{logic\}) logic
+        ty :: (\{\})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 :: (\{logic\})logic
+        ty :: (\{\})term
+\end{ttbox}
+leads to an error message and fails.
+\end{itemize}
+These conditions guarantee principal types~\cite{nipkow-prehofer}.
+
+\section{Precedence Grammars}
+\label{PrecedenceGrammars}
+\index{precedence grammar|(}
+
+The precise syntax of a logic is best defined by a context-free grammar.
+These grammars obey the following conventions: identifiers denote
+nonterminals, {\tt typewriter} fount denotes terminals, repetition is
+indicated by \dots, and alternatives are separated by $|$.
+
+In order to simplify the description of mathematical languages, we introduce
+an extended format which permits {\bf precedences}\index{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 q \ge p.~(A@q=\gamma) \in G
+\]
+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{PrecedenceEx}
+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 \verb$-$ binds tighter than
+\verb$*$ which binds tighter than \verb$+$, and that \verb$+$ and \verb$*$
+associate to the left and right, respectively.
+\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)$.
+
+Using these conventions and assuming $max_pri=9$, the grammar in
+Example~\ref{PrecedenceEx} 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}
+
+\index{precedence grammar|)}
+
+\section{Basic syntax *}
+
+An informal account of most of Isabelle's syntax (meta-logic, types etc) is
+contained in {\em Introduction to Isabelle}.  A precise description using a
+precedence grammar is shown in Figure~\ref{MetaLogicSyntax}.  This description
+is the basis of all extensions by object-logics.
+\begin{figure}[htb]
+\begin{center}
+\begin{tabular}{rclc}
+$prop$ &=& \ttindex{PROP} $aprop$ ~~$|$~~ {\tt(} $prop$ {\tt)} \\
+     &$|$& $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\{\}}
+                ~~$|$~~ {\tt\{} $id$ {\tt,} \dots {\tt,} $id$ {\tt\}} 
+\end{tabular}\index{*"!"!}\index{*"["|}\index{*"|"]}
+\indexbold{type@$type$}\indexbold{sort@$sort$}\indexbold{idts@$idts$}
+\indexbold{logic@$logic$}\indexbold{prop@$prop$}\indexbold{fun@$fun$}
+\end{center}
+\caption{Meta-Logic Syntax}
+\label{MetaLogicSyntax}
+\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 (see below).
+\item[$fun$] Terms potentially of function type.
+\item[$type$] Types.
+\item[$idts$] a list of identifiers, possibly constrained by types.  Note
+  that $x::nat~y$ is parsed as $x::(nat~y)$, i.e.\ $y$ is treated like a
+  type constructor applied to $nat$.
+\end{description}
+
+The predefined types $id$, $var$, $tfree$ and $tvar$ represent identifiers
+({\tt f}), unknowns ({\tt ?f}), type variables ({\tt 'a}), and type unknowns
+({\tt ?'a}) respectively.  If we think of them as nonterminals with
+predefined syntax, we may assume that all their productions have precedence
+$max_pri$.
+
+\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}
+
+
+\section{Mixfix syntax}
+\index{mixfix|(}
+
+We distinguish between abstract and concrete syntax.  The {\em abstract}
+syntax is given by the typed constants of a theory.  Abstract syntax trees are
+well-typed terms, i.e.\ values of \ML\ type {\tt term}.  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)$, where $f$ is a constant or variable.  Since this notation
+quickly becomes unreadable, Isabelle supports syntax definitions in the form
+of unrestricted context-free grammars using mixfix annotations.
+
+Mixfix annotations describe the {\em concrete} syntax, its 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 associates
+an Isabelle constant with it.
+
+A {\em mixfix declaration} {\tt consts $c$ ::\ $\tau$ ($sy$ $ps$ $p$)} is
+interpreted as a grammar pro\-duction as follows:
+\begin{itemize}
+\item $sy$ is the right-hand side of this production, specified as a {\em
+    mixfix annotation}.  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/nonterminal and the strings
+  $\alpha@i$ do not contain ``{\tt_}''.
+\item $\tau$ specifies the types of the nonterminals on the left and right
+  hand side. If $sy$ is of the form above, $\tau$ must be of the form
+  $[\tau@1,\dots,\tau@n] \To \tau'$.  Then argument $i$ is of type $\tau@i$
+  and the result, i.e.\ the left-hand side of the production, is of type
+  $\tau'$.  Both the $\tau@i$ and $\tau'$ may be function types.
+\item $c$ is the name of the Isabelle constant associated with this production.
+  Parsing an instance of the phrase $sy$ generates the {\tt term} {\tt
+    Const($c$,dummyT\footnote{Proper types are inserted later on.  See
+      \S\ref{Typing}})\$$a@1$\$$\dots$\$$a@n$}\index{*dummyT}, where $a@i$ is
+  the term generated by parsing the $i^{th}$ argument.
+\item $ps$ must be of the form $[p@1,\dots,p@n]$, where $p@i$ is the
+  minimal precedence\index{precedence} required of any phrase that may appear
+  as the $i^{th}$ argument.  The null list is interpreted as a list of 0's of
+  the appropriate length.
+\item $p$ is the precedence of this production.
+\end{itemize}
+Notice that there is a close connection between abstract and concrete syntax:
+each production has an associated constant, and types act as {\bf syntactic
+  categories} in the concrete syntax.  To emphasize this connection, we
+sometimes refer to the nonterminals on the right-hand side of a production as
+its arguments and to the nonterminal on the left-hand side as its result.
+
+The maximal legal precedence is called \ttindexbold{max_pri}, which is
+currently 1000.  If you want to ignore precedences, the safest way to do so is
+to use the annotation {\tt($sy$)}: this production puts no precedence
+constraints on any of its arguments and has maximal precedence itself, i.e.\ 
+it is always applicable and does not exclude any productions of its
+arguments.
+
+\begin{example}
+In mixfix notation the grammar in Example~\ref{PrecedenceEx} can be written
+as follows:
+\begin{ttbox}
+types exp 0
+consts "0"  ::              "exp"  ("0" 9)
+       "+"  :: "[exp,exp] => exp"  ("_ + _" [0,1] 0)
+       "*"  :: "[exp,exp] => exp"  ("_ * _" [3,2] 2)
+       "-"  ::       "exp => exp"  ("- _"   [3]   3)
+\end{ttbox}
+Parsing the string \verb!"0 + - 0 + 0"! produces the term {\tt
+  $p$\$($p$\$($m$\$$z$)\$$z$)\$$z$} where {\tt$p =$ Const("+",dummyT)},
+{\tt$m =$ Const("-",dummyT)}, and {\tt$z =$ Const("0",dummyT)}.
+\end{example}
+
+The interpretation of \ttindex{_} in a mixfix annotation is always as a {\bf
+  meta-character}\index{meta-character} which does not represent itself but
+an argument position.  The following characters are also meta-characters:
+\begin{ttbox}
+'   (   )   /
+\end{ttbox}
+Preceding any character with a quote (\verb$'$) turns it into an ordinary
+character.  Thus you can write \verb!''! if you really want a single quote.
+The purpose of the other meta-characters is explained in
+\S\ref{PrettyPrinting}.  Remember that in \ML\ strings \verb$\$ is already a
+(different kind of) meta-character.
+
+
+\subsection{Types and syntactic categories *}
+
+The precise mapping from types to syntactic categories is defined by the
+following function:
+\begin{eqnarray*}
+N(\tau@1\To\tau@2) &=& fun \\
+N((\tau@1,\dots,\tau@n)ty) &=& ty \\
+N(\alpha) &=& logic
+\end{eqnarray*}
+Only the outermost type constructor is taken into account and type variables
+can range over all logical types.  This catches some ill-typed terms (like
+$Cons(x,0)$, where $Cons :: [\alpha,\alpha list] \To \alpha list$ and $0 ::
+nat$) but leaves the real work to the type checker.
+
+In terms of the precedence grammar format introduced in
+\S\ref{PrecedenceGrammars}, the declaration
+\begin{ttbox}
+consts \(c\) :: "[\(\tau@1\),\dots,\(\tau@n\)]\(\To\tau\)" ("\(\alpha@0\_\alpha@1\dots\alpha@{n-1}\_\alpha@n\)") [\(p@1\),\dots,\(p@n\)] \(p\))
+\end{ttbox}
+defines the production
+\[ N(\tau)@p ~~=~~ \alpha@0 ~N(\tau@1)@{p@1}~ \alpha@1~ \dots
+                  ~\alpha@{n-1} ~N(\tau@n)@{p@n}~ \alpha@n
+\]
+
+\subsection{Copy productions *}
+
+Productions which do not create a new node in the abstract syntax tree are
+called {\bf copy productions}.  They must have exactly one nonterminal on
+the right hand side.  The term generated when parsing that nonterminal is
+simply passed up as the result of parsing the whole copy production.  In
+Isabelle a copy production is indicated by an empty constant name, i.e.\ by
+\begin{ttbox}
+consts "" :: \(\tau\)  (\(sy\) \(ps\) \(p\))
+\end{ttbox}
+
+A special kind of copy production is one where, modulo white space, $sy$ is
+{\tt"_"}.  It is called a {\bf 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.  The following example demonstrates the effect:
+the grammar defined by
+\begin{ttbox}
+types A,B,C 0
+consts AB :: "B => A"  ("A _" [10] 517)
+       "" :: "C => B"  ("_"   [0]  100)
+       x  :: "C"       ("x"          5)
+       y  :: "C"       ("y"         15)
+\end{ttbox}
+admits {\tt"A y"} but not {\tt"A x"}.  Had the constant in the second
+production been some non-empty string, both {\tt"A y"} and {\tt"A x"} would
+be legal.
+
+\index{mixfix|)}
+
+\section{Lexical conventions}
+
+The lexical analyzer distinguishes the following kinds of tokens: delimiters,
+identifiers, unknowns, type variables and type unknowns.
+
+Delimiters are user-defined, i.e.\ they are extracted from the syntax
+definition.  If $\alpha@0\_\alpha@1\dots\alpha@{n-1}\_\alpha@n$ is a mixfix
+annotation, each $\alpha@i$ is decomposed into substrings
+$\beta@1~\dots~\beta@k$ which are separated by and do not contain
+\bfindex{white space} ( = blanks, tabs, newlines).  Each $\beta@j$ becomes a
+delimiter.  Thus a delimiter can be an arbitrary string not containing white
+space.
+
+The lexical syntax of identifiers and variables ( = unknowns) is defined in
+the introductory manual.  Parsing an identifier $f$ generates {\tt
+  Free($f$,dummyT)}\index{*dummyT}.  Parsing a variable {\tt?}$v$ generates
+{\tt Var(($u$,$i$),dummyT)} where $i$ is the integer value of the longest
+numeric suffix of $v$ (possibly $0$), and $u$ is the remaining prefix.
+Parsing a variable {\tt?}$v{.}i$ generates {\tt Var(($v$,$i$),dummyT)}.  The
+following table covers the four different cases that can arise:
+\begin{center}\tt
+\begin{tabular}{cccc}
+"?v" & "?v.7" & "?v5" & "?v7.5" \\
+Var(("v",0),$d$) & Var(("v",7),$d$) & Var(("v",5),$d$) & Var(("v7",5),$d$)
+\end{tabular}
+\end{center}
+where $d = {\tt dummyT}$.
+
+In mixfix annotations, \ttindexbold{id}, \ttindexbold{var},
+\ttindexbold{tfree} and \ttindexbold{tvar} are the predefined categories of
+identifiers, unknowns, type variables and type unknowns, respectively.
+
+
+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 delimiter and an identifier or variable (like
+{\tt ALL}) is treated as a delimiter.  White spaces are separators.
+
+An important consequence of this translation scheme is that delimiters need
+not be separated by white space to be recognized as separate.  If \verb$"-"$
+is a delimiter but \verb$"--"$ is not, the string \verb$"--"$ is treated as
+two consecutive occurrences of \verb$"-"$.  This is in contrast to \ML\ which
+would treat \verb$"--"$ as a single (undeclared) identifier.  The
+consequence of Isabelle's more liberal scheme is that the same string may be
+parsed in a different way after extending the syntax: after adding
+\verb$"--"$ as a delimiter, the input \verb$"--"$ is treated as
+a single occurrence of \verb$"--"$.
+
+\section{Infix operators}
+
+{\tt Infixl} and {\tt infixr} declare infix operators which associate to the
+left and right respectively.  As in \ML, prefixing infix operators with
+\ttindexbold{op} turns them into curried functions.  Infix declarations can
+be reduced to mixfix ones as follows:
+\begin{center}\tt
+\begin{tabular}{l@{~~$\Longrightarrow$~~}l}
+"$c$" ::~$\tau$ (\ttindexbold{infixl} $p$) &
+"op $c$" ::~$\tau$ ("_ $c$ _" [$p$,$p+1$] $p$) \\
+"$c$" ::~$\tau$ (\ttindexbold{infixr} $p$) &
+"op $c$" ::~$\tau$ ("_ $c$ _" [$p+1$,$p$] $p$)
+\end{tabular}
+\end{center}
+
+
+\section{Binders}
+A {\bf binder} is a variable-binding constant, such as a quantifier.
+The declaration
+\begin{ttbox}
+consts \(c\) :: \(\tau\)  (binder \(Q\) \(p\))
+\end{ttbox}\indexbold{*binder}
+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.t$.  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 $t$, and
+$\tau@3$ the type of the whole term.  For example $\forall$ can be declared
+like this:
+\begin{ttbox}
+consts 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$)}; the latter form is for purists only.
+
+In case $\tau@2 = \tau@3$, nested quantifications can be written as $Q~x@1
+\dots x@n.t$.  From a syntactic point of view,
+\begin{ttbox}
+consts \(c\) :: "\((\tau@1\To\tau@2)\To\tau@3\)"  (binder "\(Q\)" \(p\))
+\end{ttbox}
+is equivalent to
+\begin{ttbox}
+consts \(c\)   :: "\((\tau@1\To\tau@2)\To\tau@3\)"
+       "\(Q\)" :: "[idts,\(\tau@2\)] => \(\tau@3\)"  ("\(Q\)_.  _" \(p\))
+\end{ttbox}
+where {\tt idts} is the syntactic category $idts$ defined in
+Figure~\ref{MetaLogicSyntax}.
+
+However, there is more to binders than concrete syntax: behind the scenes the
+body of the quantified expression has to be converted into a
+$\lambda$-abstraction (when parsing) and back again (when printing).  This
+is performed by the translation mechanism, which is discussed below.  For
+binders, the definition of the required translation functions has been
+automated.  Many other syntactic forms, such as set comprehension, require
+special treatment.
+
+
+\section{Parse translations *}
+\label{Parse-translations}
+\index{parse translation|(}
+
+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 each production.  In many
+cases this scheme is not powerful enough, especially for constructs involving
+variable bindings.  Therefore the $ML$-section of a theory definition can
+associate constant names with user-defined translation functions by including
+a line
+\begin{ttbox}
+val parse_translation = \dots
+\end{ttbox}
+where the right-hand side of this binding must be an \ML-expression of type
+\verb$(string * (term list -> term))list$.
+
+After the input string has been translated into a term according to the
+syntax definition, there is a second phase in which the term is translated
+using the user-supplied functions in a bottom-up manner.  Given a list $tab$
+of the above type, a term $t$ is translated as follows.  If $t$ is not of the
+form {\tt Const($c$,$\tau$)\$$t@1$\$\dots\$$t@n$}, then $t$ is returned
+unchanged.  Otherwise all $t@i$ are translated into $t@i'$.  Let {\tt $t' =$
+  Const($c$,$\tau$)\$$t@1'$\$\dots\$$t@n'$}.  If there is no pair $(c,f)$ in
+$tab$, return $t'$.  Otherwise apply $f$ to $[t@1',\dots,t@n']$.  If that
+raises an exception, return $t'$, otherwise return the result.
+\begin{example}\label{list-enum}
+\ML-lists are constructed by {\tt[]} and {\tt::}.  For readability the
+list \hbox{\tt$x$::$y$::$z$::[]} can be written \hbox{\tt[$x$,$y$,$z$]}.
+In Isabelle the two forms of lists are declared as follows:
+\begin{ttbox}
+types list 1
+      enum 0
+arities list :: (term)term
+consts "[]" :: "'a list"                   ("[]")
+       ":"  :: "['a, 'a list] => 'a list"  (infixr 50)
+       enum :: "enum => 'a list"           ("[_]")
+       sing :: "'a => enum"                ("_")
+       cons :: "['a,enum] => enum"         ("_, _")
+end
+\end{ttbox}
+Because \verb$::$ is already used for type constraints, it is replaced by
+\verb$:$ as the infix list constructor.
+
+In order to allow list enumeration, the new type {\tt enum} is introduced.
+Its only purpose is syntactic and hence it does not need an arity, in
+contrast to the logical type {\tt list}.  Although \hbox{\tt[$x$,$y$,$z$]} is
+syntactically legal, it needs to be translated into a term built up from
+\verb$[]$ and \verb$:$.  This is what \verb$make_list$ accomplishes:
+\begin{ttbox}
+val cons = Const("op :", dummyT);
+
+fun make_list (Const("sing",_)$e) = cons $ e $ Const("[]", dummyT)
+  | make_list (Const("cons",_)$e$es) = cons $ e $ make_list es;
+\end{ttbox}
+To hook this translation up to Isabelle's parser, the theory definition needs
+to contain the following $ML$-section:
+\begin{ttbox}
+ML
+fun enum_tr[enum] = make_list enum;
+val parse_translation = [("enum",enum_tr)]
+\end{ttbox}
+This causes \verb!Const("enum",_)$!$t$ to be replaced by
+\verb$enum_tr[$$t$\verb$]$.
+
+Of course the definition of \verb$make_list$ should be included in the
+$ML$-section.
+\end{example}
+\begin{example}\label{SET}
+  Isabelle represents the set $\{ x \mid P(x) \}$ internally by $Set(\lambda
+  x.P(x))$.  The internal and external forms need separate
+constants:\footnote{In practice, the external form typically has a name
+beginning with an {\at} sign, such as {\tt {\at}SET}.  This emphasizes that
+the constant should be used only for parsing/printing.}
+\begin{ttbox}
+types set 1
+arities set :: (term)term
+consts Set :: "('a => o) => 'a set"
+       SET :: "[id,o] => 'a set"  ("\{_ | _\}")
+\end{ttbox}
+Parsing {\tt"\{$x$ | $P$\}"} according to this syntax yields the term {\tt
+  Const("SET",dummyT) \$ Free("\(x\)",dummyT) \$ \(p\)}, where $p$ is the
+result of parsing $P$.  What we need is the term {\tt
+  Const("Set",dummyT)\$Abs("$x$",dummyT,$p'$)}, where $p'$ is some
+``abstracted'' version of $p$.  Therefore we define a function
+\begin{ttbox}
+fun set_tr[Free(s,T), p] = Const("Set", dummyT) $
+                           Abs(s, T, abstract_over(Free(s,T), p));
+\end{ttbox}
+where \verb$abstract_over: term*term -> term$ is a predefined function such
+that {\tt abstract_over($u$,$t$)} replaces every occurrence of $u$ in $t$ by
+a {\tt Bound} variable of the correct index (i.e.\ 0 at top level).  Remember
+that {\tt dummyT} is replaced by the correct types at a later stage (see
+\S\ref{Typing}).  Function {\tt set_tr} is associated with {\tt SET} by
+including the \ML-text
+\begin{ttbox}
+val parse_translation = [("SET", set_tr)];
+\end{ttbox}
+\end{example}
+
+If you want to run the above examples in Isabelle, you should note that an
+$ML$-section needs to contain not just a definition of
+\verb$parse_translation$ but also of a variable \verb$print_translation$.  The
+purpose of the latter is to reverse the effect of the former during printing;
+details are found in \S\ref{Print-translations}.  Hence you need to include
+the line
+\begin{ttbox}
+val print_translation = [];
+\end{ttbox}
+This is instructive because the terms are then printed out in their internal
+form.  For example the input \hbox{\tt[$x$,$y$,$z$]} is echoed as
+\hbox{\tt$x$:$y$:$z$:[]}.  This helps to check that your parse translation is
+working correctly.
+
+%\begin{warn}
+%Explicit type constraints disappear with type checking but are still
+%visible to the parse translation functions.
+%\end{warn}
+
+\index{parse translation|)}
+
+\section{Printing}
+
+Syntax definitions provide printing information in three distinct ways:
+through
+\begin{itemize}
+\item the syntax of the language (as used for parsing),
+\item pretty printing information, and
+\item print translation functions.
+\end{itemize}
+The bare mixfix declarations enable Isabelle to print terms, but the result
+will not necessarily be pretty and may look different from what you expected.
+To produce a pleasing layout, you need to read the following sections.
+
+\subsection{Printing with mixfix declarations}
+
+Let {\tt$t =$ Const($c$,_)\$$t@1$\$\dots\$$t@n$} be a term and let
+\begin{ttbox}
+consts \(c\) :: \(\tau\) (\(sy\))
+\end{ttbox}
+be a mixfix declaration where $sy$ is of the form
+$\alpha@0\_\alpha@1\dots\alpha@{n-1}\_\alpha@n$.  Printing $t$ according to
+$sy$ means printing the string
+$\alpha@0\beta@1\alpha@1\ldots\alpha@{n-1}\beta@n\alpha@n$, where $\beta@i$
+is the result of printing $t@i$.
+
+Note that the system does {\em not\/} insert blanks.  They should be part of
+the mixfix syntax if they are required to separate tokens or achieve a
+certain layout.
+
+\subsection{Pretty printing}
+\label{PrettyPrinting}
+\index{pretty printing}
+
+In order to format the output, it is possible to embed pretty printing
+directives in mixfix annotations.  These directives are ignored during parsing
+and affect only printing.  The characters {\tt(}, {\tt)} and {\tt/} are
+interpreted as meta-characters\index{meta-character} when found in a mixfix
+annotation.  Their meaning is
+\begin{description}
+\item[~{\tt(}~] Open a block.  A sequence of digits following it is
+  interpreted as the \bfindex{indentation} of this block.  It causes the
+  output to be indented by $n$ positions if a line break occurs within the
+  block.  If {\tt(} is not followed by a digit, the indentation defaults to
+  $0$.
+\item[~{\tt)}~] Close a block.
+\item[~\ttindex{/}~] Allow a \bfindex{line break}.  White space immediately
+  following {\tt/} is not printed if the line is broken at this point.
+\end{description}
+
+\subsection{Print translations *}
+\label{Print-translations}
+\index{print translation|(}
+
+Since terms are translated after parsing (see \S\ref{Parse-translations}),
+there is a similar mechanism to translate them back before printing.
+Therefore the $ML$-section of a theory definition can associate constant
+names with user-defined translation functions by including a line
+\begin{ttbox}
+val print_translation = \dots
+\end{ttbox}
+where the right-hand side of this binding is again an \ML-expression of type
+\verb$(string * (term list -> term))list$.
+Including a pair $(c,f)$ in this list causes the printer to print
+$f[t@1,\dots,t@n]$ whenever it finds {\tt Const($c$,_)\$$t@1$\$\dots\$$t@n$}.
+\begin{example}
+Reversing the effect of the parse translation in Example~\ref{list-enum} is
+accomplished by the following function:
+\begin{ttbox}
+fun make_enum (Const("op :",_) $ e $ es) = case es of
+        Const("[]",_)  =>  Const("sing",dummyT) $ e
+      | _  =>  Const("enum",dummyT) $ e $ make_enum es;
+\end{ttbox}
+It translates \hbox{\tt$x$:$y$:$z$:[]} to \hbox{\tt[$x$,$y$,$z$]}.  However,
+if the input does not terminate with an empty list, e.g.\ \hbox{\tt$x$:$xs$},
+\verb$make_enum$ raises exception {\tt Match}.  This signals that the
+attempted translation has failed and the term should be printed as is.
+The connection with Isabelle's pretty printer is established as follows:
+\begin{ttbox}
+fun enum_tr'[x,xs] = Const("enum",dummyT) $
+                     make_enum(Const("op :",dummyT)$x$xs);
+val print_translation = [("op :", enum_tr')];
+\end{ttbox}
+This declaration causes the printer to print \hbox{\tt enum_tr'[$x$,$y$]}
+whenever it finds \verb!Const("op :",_)$!$x$\verb!$!$y$.
+\end{example}
+\begin{example}
+  In Example~\ref{SET} we showed how to translate the concrete syntax for set
+  comprehension into the proper internal form.  The string {\tt"\{$x$ |
+    $P$\}"} now becomes {\tt Const("Set",_)\$Abs("$x$",_,$p$)}.  If, however,
+  the latter term were printed without translating it back, it would result
+  in {\tt"Set(\%$x$.$P$)"}.  Therefore the abstraction has to be turned back
+  into a term that matches the concrete mixfix syntax:
+\begin{ttbox}
+fun set_tr'[Abs(x,T,P)] =
+      let val (x',P') = variant_abs(x,T,P)
+      in Const("SET",dummyT) $ Free(x',T) $ P' end;
+\end{ttbox}
+The function \verb$variant_abs$, a basic term manipulation function, replaces
+the bound variable $x$ by a {\tt Free} variable $x'$ having a unique name.  A
+term produced by {\tt set_tr'} can now be printed according to the concrete
+syntax defined in Example~\ref{SET} above.
+
+Notice that the application of {\tt set_tr'} fails if the second component of
+the argument is not an abstraction, but for example just a {\tt Free}
+variable.  This is intentional because it signals to the caller that the
+translation is inapplicable.  As a result {\tt Const("Set",_)\$Free("P",_)}
+prints as {\tt"Set(P)"}.
+
+The full theory extension, including concrete syntax and both translation
+functions, has the following form:
+\begin{ttbox}
+types set 1
+arities set :: (term)term
+consts Set :: "('a => o) => 'a set"
+       SET :: "[id,o] => 'a set"  ("\{_ | _\}")
+end
+ML
+fun set_tr[Free(s,T), p] = \dots;
+val parse_translation = [("SET", set_tr)];
+fun set_tr'[Abs(x,T,P)] = \dots;
+val print_translation = [("Set", set_tr')];
+\end{ttbox}
+\end{example}
+As during the parse translation process, the types attached to constants
+during print translation are ignored.  Thus {\tt Const("SET",dummyT)} in
+{\tt set_tr'} above is acceptable.  The types of {\tt Free}s and {\tt Var}s
+however must be preserved because they may get printed (see {\tt
+show_types}).
+
+\index{print translation|)}
+
+\subsection{Printing a term}
+
+Let $tab$ be the set of all string-function pairs of print translations in the
+current syntax.
+
+Terms are printed recursively; print translations are applied top down:
+\begin{itemize}
+\item {\tt Free($x$,_)} is printed as $x$.
+\item {\tt Var(($x$,$i$),_)} is printed as $x$, if $i = 0$ and $x$ does not
+  end with a digit, as $x$ followed by $i$, if $i \neq 0$ and $x$ does not
+  end with a digit, and as {\tt $x$.$i$}, if $x$ ends with a digit.  Thus the
+  following cases can arise:
+\begin{center}
+{\tt\begin{tabular}{cccc}
+\verb$Var(("v",0),_)$ & \verb$Var(("v",7),_)$ & \verb$Var(("v5",0),_)$ \\
+"?v" & "?v7" & "?v5.0"
+\end{tabular}}
+\end{center}
+\item {\tt Abs($x@1$,_,Abs($x@2$,_,\dots Abs($x@n$,_,$p$)\dots))}, where $p$
+  is not an abstraction, is printed as {\tt \%$y@1\dots y@n$.$P$}, where $P$
+  is the result of printing $p$, and the $x@i$ are replaced by $y@i$.  The
+  latter are (possibly new) unique names.
+\item {\tt Bound($i$)} is printed as {\tt B.$i$} \footnote{The occurrence of
+    such ``loose'' bound variables indicates that either you are trying to
+    print a subterm of an abstraction, or there is something wrong with your
+    print translations.}.
+\item The application {\tt$t =$ Const($c$,_)\$$t@1$\$\dots\$$t@n$} (where
+  $n$ may be $0$!) is printed as follows:
+
+  If there is a pair $(c,f)$ in $tab$, print $f[t@1,\dots,t@n]$.  If this
+  application raises exception {\tt Match} or there is no pair $(c,f)$ in
+  $tab$, let $sy$ be the mixfix annotation associated with $c$.  If there is
+  no such $sy$, or if $sy$ does not have exactly $n$ argument positions, $t$
+  is printed as an application; otherwise $t$ is printed according to $sy$.
+
+  All other applications are printed as applications.
+\end{itemize}
+Printing a term {\tt $c$\$$t@1$\$\dots\$$t@n$} as an application means
+printing it as {\tt $s$($s@1$,\dots,$s@n$)}, where $s@i$ is the result of
+printing $t@i$.  If $c$ is a {\tt Const}, $s$ is its first argument;
+otherwise $s$ is the result of printing $c$ as described above.
+\medskip
+
+The printer also inserts parentheses where they are necessary for reasons
+of precedence.
+
+\section{Identifiers, constants, and type inference *}
+\label{Typing}
+
+There is one final step in the translation from strings to terms that we have
+not covered yet.  It explains how constants are distinguished from {\tt Free}s
+and how {\tt Free}s and {\tt Var}s are typed.  Both issues arise because {\tt
+  Free}s and {\tt Var}s are not declared.
+
+An identifier $f$ that does not appear as a delimiter in the concrete syntax
+can be either a free variable or a constant.  Since the parser knows only
+about those constants which appear in mixfix annotations, it parses $f$ as
+{\tt Free("$f$",dummyT)}, where \ttindex{dummyT} is the predefined dummy {\tt
+  typ}.  Although the parser produces these very raw terms, most user
+interface level functions like {\tt goal} type terms according to the given
+theory, say $T$.  In a first step, every occurrence of {\tt Free($f$,_)} or
+{\tt Const($f$,_)} is replaced by {\tt Const($f$,$\tau$)}, provided there is
+a constant $f$ of {\tt typ} $\tau$ in $T$.  This means that identifiers are
+treated as {\tt Free}s iff they are not declared in the theory.  The types of
+the remaining {\tt Free}s (and {\tt Var}s) are inferred as in \ML.  Type
+constraints can be used to remove ambiguities.
+
+One peculiarity of the current type inference algorithm is that variables
+with the same name must have the same type, irrespective of whether they are
+schematic, free or bound.  For example, take the first-order formula $f(x) = x
+\land (\forall f.~ f=f)$ where ${=} :: [\alpha{::}term,\alpha]\To o$ and
+$\forall :: (\alpha{::}term\To o)\To o$.  The first conjunct forces
+$x::\alpha{::}term$ and $f::\alpha\To\alpha$, the second one
+$f::\beta{::}term$.  Although the two $f$'s are distinct, they are required to
+have the same type.  Unifying $\alpha\To\alpha$ and $\beta{::}term$ fails
+because, in first-order logic, function types are not in class $term$.
+
+
+\section{Putting it all together}
+
+Having discussed the individual building blocks of a logic definition, it
+remains to be shown how they fit together.  In particular we need to say how
+an object-logic syntax is hooked up to the meta-logic.  Since all theorems
+must conform to the syntax for $prop$ (see Figure~\ref{MetaLogicSyntax}),
+that syntax has to be extended with the object-level syntax.  Assume that the
+syntax of your object-logic defines a category $o$ of formulae.  These
+formulae can now appear in axioms and theorems wherever $prop$ does if you
+add the production
+\[ prop ~=~ form.  \]
+More precisely, you need a coercion from formulae to propositions:
+\begin{ttbox}
+Base = Pure +
+types o 0
+arities o :: logic
+consts Trueprop :: "o => prop"  ("_"  5)
+end
+\end{ttbox}
+The constant {\tt Trueprop} (the name is arbitrary) acts as an invisible
+coercion function.  Assuming this definition resides in a file {\tt base.thy},
+you have to load it with the command {\tt use_thy"base"}.
+
+One of the simplest nontrivial logics is {\em minimal logic} of
+implication.  Its definition in Isabelle needs no advanced features but
+illustrates the overall mechanism quite nicely:
+\begin{ttbox}
+Hilbert = Base +
+consts "-->" :: "[o,o] => o"  (infixr 10)
+rules
+K   "P --> Q --> P"
+S   "(P --> Q --> R) --> (P --> Q) --> P --> R"
+MP  "[| P --> Q; P |] ==> Q"
+end
+\end{ttbox}
+After loading this definition you can start to prove theorems in this logic:
+\begin{ttbox}
+goal Hilbert.thy "P --> P";
+{\out Level 0}
+{\out P --> P}
+{\out  1.  P --> P}
+by (resolve_tac [Hilbert.MP] 1);
+{\out Level 1}
+{\out P --> P}
+{\out  1.  ?P --> P --> P}
+{\out  2.  ?P}
+by (resolve_tac [Hilbert.MP] 1);
+{\out Level 2}
+{\out P --> P}
+{\out  1.  ?P1 --> ?P --> P --> P}
+{\out  2.  ?P1}
+{\out  3.  ?P}
+by (resolve_tac [Hilbert.S] 1);
+{\out Level 3}
+{\out P --> P}
+{\out  1.  P --> ?Q2 --> P}
+{\out  2.  P --> ?Q2}
+by (resolve_tac [Hilbert.K] 1);
+{\out Level 4}
+{\out P --> P}
+{\out  1.  P --> ?Q2}
+by (resolve_tac [Hilbert.K] 1);
+{\out Level 5}
+{\out P --> P}
+{\out No subgoals!}
+\end{ttbox}
+As you can see, this Hilbert-style formulation of minimal logic is easy to
+define but difficult to use.  The following natural deduction formulation is
+far preferable:
+\begin{ttbox}
+MinI = Base +
+consts "-->" :: "[o,o] => o"  (infixr 10)
+rules
+impI  "(P ==> Q) ==> P --> Q"
+impE  "[| P --> Q; P |] ==> Q"
+end
+\end{ttbox}
+Note, however, that although the two systems are equivalent, this fact cannot
+be proved within Isabelle: {\tt S} and {\tt K} can be derived in \verb$MinI$
+(exercise!), but {\tt impI} cannot be derived in \verb!Hilbert!.  The reason
+is that {\tt impI} is only an {\em admissible} rule in \verb!Hilbert!,
+something that can only be shown by induction over all possible proofs in
+\verb!Hilbert!.
+
+It is a very simple matter to extend minimal logic with falsity:
+\begin{ttbox}
+MinIF = MinI +
+consts False :: "o"
+rules
+FalseE  "False ==> P"
+end
+\end{ttbox}
+On the other hand, we may wish to introduce conjunction only:
+\begin{ttbox}
+MinC = Base +
+consts "&" :: "[o,o] => o"  (infixr 30)
+rules
+conjI  "[| P; Q |] ==> P & Q"
+conjE1 "P & Q ==> P"
+conjE2 "P & Q ==> Q"
+end
+\end{ttbox}
+And if we want to have all three connectives together, we define:
+\begin{ttbox}
+MinIFC = MinIF + MinC
+\end{ttbox}
+Now we can prove mixed theorems like
+\begin{ttbox}
+goal MinIFC.thy "P & False --> Q";
+by (resolve_tac [MinI.impI] 1);
+by (dresolve_tac [MinC.conjE2] 1);
+by (eresolve_tac [MinIF.FalseE] 1);
+\end{ttbox}
+Try this as an exercise!
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/abstract.txt	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,8 @@
+Isabelle Reference Manual.  Report 283
+
+This manual is a comprehensive description of Isabelle, including all
+commands, functions and packages.  Functions are organized according to the
+task they perform.  In each section, basic functions appear before advanced
+ones.  The Index provides an alphabetical listing.  It is intended as a
+reference, not for casual reading.  The manual assumes familiarity with the
+basic concepts explained in Report 280, Introduction to Isabelle.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/classical.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,408 @@
+%% $Id$
+\chapter{The classical theorem prover}
+\index{classical prover|(}
+\newcommand\ainfer[2]{\begin{array}{r@{}l}#2\\ \hline#1\end{array}}
+Although Isabelle is generic, many users will be working in some extension
+of classical first-order logic (FOL).  Isabelle's set theory is built upon
+FOL, while higher-order logic contains FOL as a fragment.  Theorem-proving
+in FOL is of course undecidable, but many researchers have developed
+strategies to assist in this task.
+
+Isabelle's classical module is an \ML{} functor that accepts certain
+information about a logic and delivers a suite of automatic tactics.  Each
+tactic takes a collection of rules and executes a simple, non-clausal proof
+procedure.  They are slow and simplistic compared with resolution theorem
+provers, but they can save considerable time and effort.  They can prove
+theorems such as Pelletier's~\cite{pelletier86} problems~40 and~41 in
+seconds:
+\[ (\exists y. \forall x. J(y,x) \bimp \neg J(x,x))  
+   \imp  \neg (\forall x. \exists y. \forall z. J(z,y) \bimp \neg J(z,x)) \]
+\[ (\forall z. \exists y. \forall x. F(x,y) \bimp F(x,z) \conj \neg F(x,x))
+   \imp \neg (\exists z. \forall x. F(x,z))  
+\]
+The tactics are generic.  They are not restricted to~FOL, and have been
+heavily used in the development of Isabelle's set theory.  Few interactive
+proof assistants provide this much automation.  Isabelle does not record
+proofs, but the tactics can be traced, and their components can be called
+directly.  In this manner, any proof can be viewed interactively.
+
+The logics FOL, HOL and ZF have the classical prover already installed.  We
+shall first consider how to use it, then see how to instantiate it for new
+logics. 
+
+
+\section{The sequent calculus}
+\index{sequent calculus}
+Isabelle supports natural deduction, which is easy to use for interactive
+proof.  But natural deduction does not easily lend itself to automation,
+and has a bias towards intuitionism.  For certain proofs in classical
+logic, it can not be called natural.  The {\bf sequent calculus}, a
+generalization of natural deduction, is easier to automate.
+
+A {\bf sequent} has the form $\Gamma\turn\Delta$, where $\Gamma$
+and~$\Delta$ are sets of formulae.\footnote{For FOL, sequents can
+equivalently be made from lists or multisets of formulae.}  The sequent
+\[ P@1,\ldots,P@m\turn Q@1,\ldots,Q@n \]
+is {\bf valid} if $P@1\conj\ldots\conj P@m$ implies $Q@1\disj\ldots\disj
+Q@n$.  Thus $P@1,\ldots,P@m$ represent assumptions, each of which is true,
+while $Q@1,\ldots,Q@n$ represent alternative goals.  A sequent is {\bf
+basic} if its left and right sides have a common formula, as in $P,Q\turn
+Q,R$; basic sequents are trivially valid.
+
+Sequent rules are classified as {\bf right} or {\bf left}, indicating which
+side of the $\turn$~symbol they operate on.  Rules that operate on the
+right side are analogous to natural deduction's introduction rules, and
+left rules are analogous to elimination rules.  The sequent calculus
+analogue of~$({\imp}I)$ is the rule
+$$ \ainfer{\Gamma &\turn \Delta, P\imp Q}{P,\Gamma &\turn \Delta,Q}
+   \eqno({\imp}R) $$
+This breaks down some implication on the right side of a sequent; $\Gamma$
+and $\Delta$ stand for the sets of formulae that are unaffected by the
+inference.  The analogue of the pair~$({\disj}I1)$ and~$({\disj}I2)$ is the
+single rule 
+$$ \ainfer{\Gamma &\turn \Delta, P\disj Q}{\Gamma &\turn \Delta,P,Q}
+   \eqno({\disj}R) $$
+This breaks down some disjunction on the right side, replacing it by both
+disjuncts.  Thus, the sequent calculus is a kind of multiple-conclusion logic.
+
+To illustrate the use of multiple formulae on the right, let us prove
+the classical theorem $(P\imp Q)\disj(Q\imp P)$.  Working backwards, we
+reduce this formula to a basic sequent:
+\[ \infer[(\disj)R]{\turn(P\imp Q)\disj(Q\imp P)}
+   {\infer[(\imp)R]{\turn(P\imp Q), (Q\imp P)\;}
+    {\infer[(\imp)R]{P \turn Q, (Q\imp P)\qquad}
+                    {P, Q \turn Q, P\qquad\qquad}}}
+\]
+This example is typical of the sequent calculus: start with the desired
+theorem and apply rules backwards in a fairly arbitrary manner.  This yields a
+surprisingly effective proof procedure.  Quantifiers add few complications,
+since Isabelle handles parameters and schematic variables.  See Chapter~10
+of {\em ML for the Working Programmer}~\cite{paulson91} for further
+discussion.
+
+
+\section{Simulating sequents by natural deduction}
+Isabelle can represent sequents directly, as in the object-logic~{\sc lk}.
+But natural deduction is easier to work with, and most object-logics employ
+it.  Fortunately, we can simulate the sequent $P@1,\ldots,P@m\turn
+Q@1,\ldots,Q@n$ by the Isabelle formula
+\[ \List{P@1;\ldots;P@m; \neg Q@2;\ldots; \neg Q@n}\Imp Q@1, \]
+where the order of the assumptions and the choice of~$Q@1$ are arbitrary.
+Elim-resolution plays a key role in simulating sequent proofs.
+
+We can easily handle reasoning on the left.
+As discussed in the {\em Introduction to Isabelle},
+elim-resolution with the rules $(\disj E)$, $(\bot E)$ and $(\exists E)$
+achieves a similar effect as the corresponding sequent rules.  For the
+other connectives, we use sequent-style elimination rules instead of
+destruction rules.  But note that the rule $(\neg L)$ has no effect
+under our representation of sequents!
+$$ \ainfer{\neg P,\Gamma &\turn \Delta}{\Gamma &\turn \Delta,P}
+   \eqno({\neg}L) $$
+What about reasoning on the right?  Introduction rules can only affect the
+formula in the conclusion, namely~$Q@1$.  The other right-side formula are
+represented as negated assumptions, $\neg Q@2$, \ldots,~$\neg Q@n$.  In
+order to operate on one of these, it must first be exchanged with~$Q@1$.
+Elim-resolution with the {\bf swap} rule has this effect:
+$$ \List{\neg P; \; \neg R\Imp P} \Imp R   \eqno(swap)$$
+To ensure that swaps occur only when necessary, each introduction rule is
+converted into a swapped form: it is resolved with the second premise
+of~$(swap)$.  The swapped form of~$({\conj}I)$, which might be
+called~$({\neg\conj}E)$, is
+\[ \List{\neg(P\conj Q); \; \neg R\Imp P; \; \neg R\Imp Q} \Imp R. \]
+Similarly, the swapped form of~$({\imp}I)$ is
+\[ \List{\neg(P\imp Q); \; \List{\neg R;P}\Imp Q} \Imp R  \]
+Swapped introduction rules are applied using elim-resolution, which deletes
+the negated formula.  Our representation of sequents also requires the use
+of ordinary introduction rules.  If we had no regard for readability, we
+could treat the right side more uniformly by representing sequents as
+\[ \List{P@1;\ldots;P@m; \neg Q@1;\ldots; \neg Q@n}\Imp \bot. \]
+
+
+\section{Extra rules for the sequent calculus}
+As mentioned, destruction rules such as $({\conj}E1,2)$ and $(\forall E)$
+must be replaced by sequent-style elimination rules.  In addition, we need
+rules to embody the classical equivalence between $P\imp Q$ and $\neg P\disj
+Q$.  The introduction rules~$({\disj}I1,2)$ are replaced by a rule that
+simulates $({\disj}R)$:
+\[ (\neg Q\Imp P) \Imp P\disj Q \]
+The destruction rule $({\imp}E)$ is replaced by
+\[ \List{P\imp Q;\; \neg P\imp R;\; Q\imp R} \Imp R. \]
+Quantifier replication also requires special rules.  In classical logic,
+$\exists x{.}P$ is equivalent to $\forall x{.}P$, and the rules $(\exists R)$
+and $(\forall L)$ are dual:
+\[ \ainfer{\Gamma &\turn \Delta, \exists x{.}P}
+          {\Gamma &\turn \Delta, \exists x{.}P, P[t/x]} \; (\exists R)
+   \qquad
+   \ainfer{\forall x{.}P, \Gamma &\turn \Delta}
+          {P[t/x], \forall x{.}P, \Gamma &\turn \Delta} \; (\forall L)
+\]
+Thus both kinds of quantifier may be replicated.  Theorems requiring
+multiple uses of a universal formula are easy to invent; consider 
+$(\forall x.P(x)\imp P(f(x))) \conj P(a) \imp P(f^n(x))$, for any~$n>1$.
+Natural examples of the multiple use of an existential formula are rare; a
+standard one is $\exists x.\forall y. P(x)\imp P(y)$.
+
+Forgoing quantifier replication loses completeness, but gains decidability,
+since the search space becomes finite.  Many useful theorems can be proved
+without replication, and the search generally delivers its verdict in a
+reasonable time.  To adopt this approach, represent the sequent rules
+$(\exists R)$, $(\exists L)$ and $(\forall R)$ by $(\exists I)$, $(\exists
+E)$ and $(\forall I)$, respectively, and put $(\forall E)$ into elimination
+form:
+$$ \List{\forall x{.}P(x); P(t)\Imp Q} \Imp Q    \eqno(\forall E@2) $$
+Elim-resolution with this rule will delete the universal formula after a
+single use.  To replicate universal quantifiers, replace the rule by
+$$ \List{\forall x{.}P(x);\; \List{P(t); \forall x{.}P(x)}\Imp Q} \Imp Q.
+   \eqno(\forall E@3) $$
+To replicate existential quantifiers, replace $(\exists I)$ by
+\[ \Bigl(\neg(\exists x{.}P(x)) \Imp P(t)\Bigr) \Imp \exists x{.}P(x). \]
+All introduction rules mentioned above are also useful in swapped form.
+
+Replication makes the search space infinite; we must apply the rules with
+care.  The classical package distinguishes between safe and unsafe
+rules, applying the latter only when there is no alternative.  Depth-first
+search may well go down a blind alley; best-first search is better behaved
+in an infinite search space.  However, quantifier replication is too
+expensive to prove any but the simplest theorems.
+
+
+\section{Classical rule sets}
+\index{classical sets|bold}
+Each automatic tactic takes a {\bf classical rule set} -- a collection of
+rules, classified as introduction or elimination and as {\bf safe} or {\bf
+unsafe}.  In general, safe rules can be attempted blindly, while unsafe
+rules must be used with care.  A safe rule must never reduce a provable
+goal to an unprovable set of subgoals.  The rule~$({\disj}I1)$ is obviously
+unsafe, since it reduces $P\disj Q$ to~$P$; fortunately, we do not require
+this rule.
+
+In general, any rule is unsafe whose premises contain new unknowns.  The
+elimination rule~$(\forall E@2)$ is unsafe, since it is applied via
+elim-resolution, which discards the assumption $\forall x{.}P(x)$ and
+replaces it by the weaker assumption~$P(\Var{t})$.  The rule $({\exists}I)$
+is unsafe for similar reasons.  The rule~$(\forall E@3)$ is unsafe in a
+different sense: since it keeps the assumption $\forall x{.}P(x)$, it is
+prone to looping.  In classical first-order logic, all rules are safe
+except those mentioned above.
+
+The safe/unsafe distinction is vague, and may be regarded merely as a way
+of giving some rules priority over others.  One could argue that
+$({\disj}E)$ is unsafe, because repeated application of it could generate
+exponentially many subgoals.  Induction rules are unsafe because inductive
+proofs are difficult to set up automatically.  Any inference is unsafe that
+instantiates an unknown in the proof state --- thus \ttindex{match_tac}
+must be used, rather than \ttindex{resolve_tac}.  Even proof by assumption
+is unsafe if it instantiates unknowns shared with other subgoals --- thus
+\ttindex{eq_assume_tac} must be used, rather than \ttindex{assume_tac}.
+
+Classical rule sets belong to the abstract type \ttindexbold{claset}, which
+supports the following operations (provided the classical prover is
+installed!):
+\begin{ttbox} 
+empty_cs : claset
+addSIs   : claset * thm list -> claset                 \hfill{\bf infix 4}
+addSEs   : claset * thm list -> claset                 \hfill{\bf infix 4}
+addSDs   : claset * thm list -> claset                 \hfill{\bf infix 4}
+addIs    : claset * thm list -> claset                 \hfill{\bf infix 4}
+addEs    : claset * thm list -> claset                 \hfill{\bf infix 4}
+addDs    : claset * thm list -> claset                 \hfill{\bf infix 4}
+print_cs : claset -> unit
+\end{ttbox}
+There are no operations for deletion from a classical set.  The add
+operations do not check for repetitions.
+\begin{description}
+\item[\ttindexbold{empty_cs}] is the empty classical set.
+
+\item[\tt $cs$ addSIs $rules$] \indexbold{*addSIs}
+adds some safe introduction~$rules$ to the classical set~$cs$.
+
+\item[\tt $cs$ addSEs $rules$] \indexbold{*addSEs}
+adds some safe elimination~$rules$ to the classical set~$cs$.
+
+\item[\tt $cs$ addSDs $rules$] \indexbold{*addSDs}
+adds some safe destruction~$rules$ to the classical set~$cs$.
+
+\item[\tt $cs$ addIs $rules$] \indexbold{*addIs}
+adds some unsafe introduction~$rules$ to the classical set~$cs$.
+
+\item[\tt $cs$ addEs $rules$] \indexbold{*addEs}
+adds some unsafe elimination~$rules$ to the classical set~$cs$.
+
+\item[\tt $cs$ addDs $rules$] \indexbold{*addDs}
+adds some unsafe destruction~$rules$ to the classical set~$cs$.
+
+\item[\ttindexbold{print_cs} $cs$] prints the rules of the classical set~$cs$.
+\end{description}
+Introduction rules are those that can be applied using ordinary resolution.
+The classical set automatically generates their swapped forms, which will
+be applied using elim-resolution.  Elimination rules are applied using
+elim-resolution.  In a classical set, rules are sorted the number of new
+subgoals they will yield, so that rules that generate the fewest subgoals
+will be tried first (see \S\ref{biresolve_tac}).
+
+For a given classical set, the proof strategy is simple.  Perform as many
+safe inferences as possible; or else, apply certain safe rules, allowing
+instantiation of unknowns; or else, apply an unsafe rule.  The tactics may
+also apply \ttindex{hyp_subst_tac}, if they have been set up to do so (see
+below).  They may perform a form of Modus Ponens: if there are assumptions
+$P\imp Q$ and~$P$, then replace $P\imp Q$ by~$Q$.
+
+
+\section{The classical tactics}
+\index{classical prover!tactics|bold}
+\index{tactics!for classical proof|bold}
+If installed, the classical module provides several tactics (and other
+operations) for simulating the classical sequent calculus.
+
+\subsection{Single-step tactics}
+\begin{ttbox} 
+safe_step_tac : claset -> int -> tactic
+safe_tac      : claset        -> tactic
+inst_step_tac : claset -> int -> tactic
+step_tac      : claset -> int -> tactic
+slow_step_tac : claset -> int -> tactic
+\end{ttbox}
+The automatic proof procedures call these tactics.  By calling them
+yourself, you can execute these procedures one step at a time.
+\begin{description}
+\item[\ttindexbold{safe_step_tac} $cs$ $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{safe_tac} $cs$] repeatedly performs safe steps on all 
+subgoals.  It is deterministic, with at most one outcome.  If the automatic
+tactics fail, try using {\tt safe_tac} to open up your formula; then you
+can replicate certain quantifiers explicitly by applying appropriate rules.
+
+\item[\ttindexbold{inst_step_tac} $cs$ $i$] is like {\tt safe_step_tac},
+but allows unknowns to be instantiated.
+
+\item[\ttindexbold{step_tac} $cs$ $i$] tries {\tt safe_tac}.  IF this
+fails, it tries {\tt inst_step_tac}, or applies an unsafe rule from~$cs$.
+This is the basic step of the proof procedure.
+
+\item[\ttindexbold{slow_step_tac}] 
+  resembles {\tt step_tac}, but allows backtracking between using safe
+  rules with instantiation ({\tt inst_step_tac}) and using unsafe rules.
+  The resulting search space is too large for use in the standard proof
+  procedures, but {\tt slow_step_tac} is worth considering in special
+  situations.
+\end{description}
+
+
+\subsection{The automatic tactics}
+\begin{ttbox} 
+fast_tac : claset -> int -> tactic
+best_tac : claset -> int -> tactic
+\end{ttbox}
+Both of these tactics work by applying {\tt step_tac} repeatedly.  Their
+effect is restricted (by {\tt SELECT_GOAL}) to one subgoal; they either
+solve this subgoal or fail.
+\begin{description}
+\item[\ttindexbold{fast_tac} $cs$ $i$] applies {\tt step_tac} using
+depth-first search, to solve subgoal~$i$.
+
+\item[\ttindexbold{best_tac} $cs$ $i$] applies {\tt step_tac} using
+best-first search, to solve subgoal~$i$.  A heuristic function ---
+typically, the total size of the proof state --- guides the search.  This
+function is supplied when the classical module is set up.
+\end{description}
+
+
+\subsection{Other useful tactics}
+\index{tactics!for contradiction|bold}
+\index{tactics!for Modus Ponens|bold}
+\begin{ttbox} 
+contr_tac    :             int -> tactic
+mp_tac       :             int -> tactic
+eq_mp_tac    :             int -> tactic
+swap_res_tac : thm list -> int -> tactic
+\end{ttbox}
+These can be used in the body of a specialized search.
+\begin{description}
+\item[\ttindexbold{contr_tac} {\it i}] solves subgoal~$i$ by detecting a
+contradiction among two assumptions of the form $P$ and~$\neg P$, or fail.
+It may instantiate unknowns.  The tactic can produce multiple outcomes,
+enumerating all possible contradictions.
+
+\item[\ttindexbold{mp_tac} {\it i}] 
+is like {\tt contr_tac}, but also attempts to perform Modus Ponens in
+subgoal~$i$.  If there are assumptions $P\imp Q$ and~$P$, then it replaces
+$P\imp Q$ by~$Q$.  It may instantiate unknowns.  It fails if it can do
+nothing.
+
+\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{swap_res_tac} {\it thms} {\it i}] refines subgoal~$i$ of
+the proof state using {\it thms}, which should be a list of introduction
+rules.  First, it attempts to solve the goal using \ttindex{assume_tac} or
+{\tt contr_tac}.  It then attempts to apply each rule in turn, attempting
+resolution and also elim-resolution with the swapped form.
+\end{description}
+
+\subsection{Creating swapped rules}
+\begin{ttbox} 
+swapify   : thm list -> thm list
+joinrules : thm list * thm list -> (bool * thm) list
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{swapify} {\it thms}] returns a list consisting of the
+swapped versions of~{\it thms}, regarded as introduction rules.
+
+\item[\ttindexbold{joinrules} \tt({\it intrs}, {\it elims})]
+joins introduction rules, their swapped versions, and elimination rules for
+use with \ttindex{biresolve_tac}.  Each rule is paired with~{\tt false}
+(indicating ordinary resolution) or~{\tt true} (indicating
+elim-resolution).
+\end{description}
+
+
+\section{Setting up the classical prover}
+\index{classical prover!setting up|bold}
+Isabelle's classical object-logics, including {\tt FOL} and {\tt HOL}, have
+the classical prover already set up.  When defining a new classical logic,
+you should set up the prover yourself.  It consists of the \ML{} functor
+\ttindex{ClassicalFun}, which takes the argument
+signature\indexbold{*CLASSICAL_DATA}
+\begin{ttbox} 
+signature CLASSICAL_DATA =
+  sig
+  val mp             : thm
+  val not_elim       : thm
+  val swap           : thm
+  val sizef          : thm -> int
+  val hyp_subst_tacs : (int -> tactic) list
+  end;
+\end{ttbox}
+Thus, the functor requires the following items:
+\begin{description}
+\item[\ttindexbold{mp}] should be the Modus Ponens rule
+$\List{\Var{P}\imp\Var{Q};\; \Var{P}} \Imp \Var{Q}$.
+
+\item[\ttindexbold{not_elim}] should be the contradiction rule
+$\List{\neg\Var{P};\; \Var{P}} \Imp \Var{R}$.
+
+\item[\ttindexbold{swap}] should be the swap rule
+$\List{\neg \Var{P}; \; \neg \Var{R}\Imp \Var{P}} \Imp \Var{R}$.
+
+\item[\ttindexbold{sizef}] is the heuristic function used for best-first
+search.  It should estimate the size of the remaining subgoals.  A good
+heuristic function is \ttindex{size_of_thm}, which measures the size of the
+proof state.  Another size function might ignore certain subgoals (say,
+those concerned with type checking).  A heuristic function might simply
+count the subgoals.
+
+\item[\ttindexbold{hyp_subst_tac}] is a list of tactics for substitution in
+the hypotheses, typically created by \ttindex{HypsubstFun} (see
+Chapter~\ref{substitution}).  This list can, of course, be empty.  The
+tactics are assumed to be safe!
+\end{description}
+The functor is not at all sensitive to the formalization of the
+object-logic.  It does not even examine the rules, but merely applies them
+according to its fixed strategy.  The functor resides in {\tt
+Provers/classical.ML} on the Isabelle distribution directory.
+
+\index{classical prover|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/goals.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,542 @@
+%% $Id$
+\chapter{Proof Management: The Subgoal Module}
+\index{proofs|(}
+\index{subgoal module|(}
+\index{reading!goals|see{proofs, starting}}
+The subgoal module stores the current proof state and many previous states;
+commands can produce new states or return to previous ones.  The {\em state
+list\/} at level $n$ is a list of pairs
+\[ [(\psi@n,\Psi@n),\; (\psi@{n-1},\Psi@{n-1}),\; \ldots,\; (\psi@0,[])] \]
+where $\psi@n$ is the current proof state, $\psi@{n-1}$ is the previous
+one, \ldots, and $\psi@0$ is the initial proof state.  The $\Psi@i$ are
+sequences (lazy lists) of proof states, storing branch points where a
+tactic returned a list longer than one.  The state lists permit various
+forms of backtracking.
+
+Chopping elements from the state list reverts to previous proof states.
+Besides this, the \ttindex{undo} command keeps a list of state lists.  The
+module actually maintains a stack of state lists, to support several
+proofs at the same time.
+
+The subgoal module always contains some proof state.  At the start of the
+Isabelle session, this state consists of a dummy formula.
+
+
+\section{Basic commands}
+Most proofs begin with {\tt goal} or {\tt goalw} and require no other
+commands than {\tt by}, {\tt chop} and {\tt undo}.  They typically end with
+a call to {\tt result}.
+\subsection{Starting a backward proof}
+\index{proofs!starting|bold}
+\begin{ttbox} 
+goal        : theory -> string -> thm list 
+goalw       : theory -> thm list -> string -> thm list 
+goalw_cterm : thm list -> Sign.cterm -> thm list 
+premises    : unit -> thm list
+\end{ttbox}
+The {\tt goal} commands start a new proof by setting the goal.  They
+replace the current state list by a new one consisting of the initial proof
+state.  They also empty the \ttindex{undo} list; this command cannot be
+undone!
+
+They all return a list of meta-hypotheses taken from the main goal.  If
+this list is non-empty, bind its value to an \ML{} identifier by typing
+something like
+\begin{ttbox} 
+val prems = it;
+\end{ttbox}
+These assumptions serve as the premises when you are deriving a rule.  They
+are also stored internally and can be retrieved later by the function {\tt
+  premises}.  When the proof is finished, \ttindex{result} compares the
+stored assumptions with the actual assumptions in the proof state.
+
+Some of the commands unfold definitions using meta-rewrite rules.  This
+expansion affects both the first subgoal and the premises, which would
+otherwise require use of \ttindex{rewrite_goals_tac} and
+\ttindex{rewrite_rule}.
+
+If the main goal has the form {\tt"!!{\it vars}.\ \ldots"}, with an outermost
+quantifier, then the list of premises will be empty.  Subgoal~1 will
+contain the meta-quantified {\it vars\/} as parameters and the goal's premises
+as assumptions.  This is useful when the next step of the proof would
+otherwise be to call \ttindex{cut_facts_tac} with the list of premises
+(\S\ref{cut_facts_tac}).
+
+\begin{description}
+\item[\ttindexbold{goal} {\it theory} {\it formula}; ] 
+begins a new proof, where {\it theory} is usually an \ML\ identifier
+and the {\it formula\/} is written as an \ML\ string.
+
+\item[\ttindexbold{goalw} {\it theory} {\it defs} {\it formula}; ] 
+is like {\tt goal} but also applies the list of {\it defs\/} as
+meta-rewrite rules to the first subgoal and the premises.
+\index{rewriting!meta-level}
+
+\item[\ttindexbold{goalw_cterm} {\it theory} {\it defs} {\it ct}; ] 
+is a version of {\tt goalw} for special applications.  The main goal is
+supplied as a cterm, not as a string.  Typically, the cterm is created from
+a term~$t$ by \hbox{\tt Sign.cterm_of (sign_of thy) $t$}.
+\index{*Sign.cterm_of}\index{*sign_of}
+
+\item[\ttindexbold{premises}()] 
+returns the list of meta-hypotheses associated with the current proof, in
+case you forget to bind them to an \ML{} identifier.
+\end{description}
+
+\subsection{Applying a tactic}
+\index{tactics!commands for applying|bold}
+\begin{ttbox} 
+by   : tactic -> unit
+byev : tactic list -> unit
+\end{ttbox}
+These commands extend the state list.  They apply a tactic to the current
+proof state.  If the tactic succeeds, it returns a non-empty sequence of
+next states.  The head of the sequence becomes the next state, while the
+tail is retained for backtracking (see~{\tt back}).
+\begin{description} \item[\ttindexbold{by} {\it tactic}; ] 
+applies the {\it tactic\/} to the proof state.
+
+\item[\ttindexbold{byev} {\it tactics}; ] 
+applies the list of {\it tactics}, one at a time.  It is useful for testing
+calls to {\tt prove_goal}, and abbreviates \hbox{\tt by (EVERY {\it
+tactics})}.
+\end{description}
+
+\noindent{\bf Error indications:}
+\begin{itemize}
+\item 
+{\tt "by:\ tactic failed"} means that the tactic returned an empty
+sequence when applied to the current proof state. 
+\item 
+{\tt "Warning:\ same as previous level"} means that the new proof state
+is identical to the previous state.
+\item
+{\tt "Warning:\ signature of proof state has changed"} means that a rule
+was applied from outside the theory of the initial proof state.  This
+guards against mistakes such as expressing the goal in intuitionistic logic
+and proving it using classical logic.
+\end{itemize}
+
+\subsection{Extracting the proved theorem}
+\index{ending a proof|bold}
+\begin{ttbox} 
+result  : unit -> thm
+uresult : unit -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{result}()] 
+returns the final theorem, after converting the free variables to
+schematics.  It discharges the assumptions supplied to the matching 
+\ttindex{goal} command.  
+
+It raises an exception unless the proof state passes certain checks.  There
+must be no assumptions other than those supplied to \ttindex{goal}.  There
+must be no subgoals.  The theorem proved must be a (first-order) instance
+of the original goal, as stated in the \ttindex{goal} command.  This allows
+{\bf answer extraction} --- instantiation of variables --- but no other
+changes to the main goal.  The theorem proved must have the same signature
+as the initial proof state.
+
+\item[\ttindexbold{uresult}()] 
+is similar but omits the checks given above.  It is needed when the initial
+goal contains function unknowns, when definitions are unfolded in the main
+goal, or when \ttindex{assume_ax} has been used.
+\end{description}
+
+
+\subsection{Undoing and backtracking}
+\index{undoing a proof command|bold}
+\index{backtracking command|see{\tt back}}
+\begin{ttbox} 
+chop    : unit -> unit
+choplev : int -> unit
+back    : unit -> unit
+undo    : unit -> unit
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{chop}();] 
+deletes the top level of the state list, cancelling the last \ttindex{by}
+command.  It provides a limited undo facility, and the {\tt undo} command
+can cancel it.
+
+\item[\ttindexbold{choplev} {\it n};] 
+truncates the state list to level~{\it n}. 
+
+\item[\ttindexbold{back}();]
+searches the state list for a non-empty branch point, starting from the top
+level.  The first one found becomes the current proof state --- the most
+recent alternative branch is taken.
+
+\item[\ttindexbold{undo}();] 
+cancels the most recent change to the proof state by the commands \ttindex{by},
+{\tt chop}, {\tt choplev}, and~{\tt back}.  It {\bf cannot}
+cancel {\tt goal} or {\tt undo} itself.  It can be repeated to
+cancel a series of commands.
+\end{description}
+\noindent{\bf Error indications for {\tt back}:}
+\begin{itemize}
+\item
+{\tt"Warning:\ same as previous choice at this level"} means that {\tt
+back} found a non-empty branch point, but that it contained the same proof
+state as the current one.  
+\item
+{\tt "Warning:\ signature of proof state has changed"} means the signature
+of the alternative proof state differs from that of the current state.
+\item 
+{\tt "back:\ no alternatives"} means {\tt back} could find no alternative
+proof state.
+\end{itemize}
+
+\subsection{Printing the proof state}
+\index{printing!proof state|bold}
+\begin{ttbox} 
+pr    : unit -> unit
+prlev : int -> unit
+goals_limit: int ref \hfill{\bf initially 10}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{pr}();] 
+prints the current proof state.
+
+\item[\ttindexbold{prlev} {\it n};] 
+prints the proof state at level {\it n}.  This allows you to review the
+previous steps of the proof.
+
+\item[\ttindexbold{goals_limit} \tt:= {\it k};] 
+specifies~$k$ as the maximum number of subgoals to print.
+\end{description}
+
+
+\subsection{Timing}
+\index{printing!timing statistics|bold}\index{proofs!timing|bold}
+\begin{ttbox} 
+proof_timing: bool ref \hfill{\bf initially false}
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{proof_timing} \tt:= true;] 
+makes the \ttindex{by} and \ttindex{prove_goal} commands display how much
+processor time was spent.  This information is compiler-dependent.
+\end{description}
+
+
+\section{Shortcuts for applying tactics}
+\index{shortcuts!for ``{\tt by}'' commands|bold}
+These commands call \ttindex{by} with common tactics.  Their chief purpose
+is to minimise typing, although the scanning shortcuts are useful in their
+own right.  Chapter~\ref{tactics} explains the tactics themselves.
+
+\subsection{Refining a given subgoal}
+\begin{ttbox} 
+ba: int -> unit
+br: thm -> int -> unit
+be: thm -> int -> unit
+bd: thm -> int -> unit
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{ba} {\it i};] 
+performs \hbox{\tt by (assume_tac {\it i});}\index{*assume_tac}
+
+\item[\ttindexbold{br} {\it thm} {\it i};] 
+performs \hbox{\tt by (resolve_tac [{\it thm}] {\it i});}\index{*resolve_tac}
+
+\item[\ttindexbold{be} {\it thm} {\it i};] 
+performs \hbox{\tt by (eresolve_tac [{\it thm}] {\it i});}\index{*eresolve_tac}
+
+\item[\ttindexbold{bd} {\it thm} {\it i};] 
+performs \hbox{\tt by (dresolve_tac [{\it thm}] {\it i});}\index{*dresolve_tac}
+\end{description}
+
+\subsubsection{Resolution with a list of theorems}
+\begin{ttbox} 
+brs: thm list -> int -> unit
+bes: thm list -> int -> unit
+bds: thm list -> int -> unit
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{brs} {\it thms} {\it i};] 
+performs \hbox{\tt by (resolve_tac {\it thms} {\it i});}\index{*resolve_tac}
+
+\item[\ttindexbold{bes} {\it thms} {\it i};] 
+performs \hbox{\tt by (eresolve_tac {\it thms} {\it i});}\index{*eresolve_tac}
+
+\item[\ttindexbold{bds} {\it thms} {\it i};] 
+performs \hbox{\tt by (dresolve_tac {\it thms} {\it i});}\index{*dresolve_tac}
+\end{description}
+
+
+\subsection{Scanning shortcuts}
+\index{shortcuts!scanning for subgoals|bold}
+These shortcuts scan for a suitable subgoal (starting from subgoal~1).
+They refine the first subgoal for which the tactic succeeds.  Thus, they
+require less typing than {\tt br}, etc.  They display the selected
+subgoal's number; please watch this, for it may not be what you expect!
+
+\subsubsection{Proof by assumption and resolution}
+\begin{ttbox} 
+fa: unit -> unit
+fr: thm -> unit
+fe: thm -> unit
+fd: thm -> unit
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{fa}();] 
+solves some subgoal by assumption.\index{*assume_tac}
+
+\item[\ttindexbold{fr} {\it thm};] 
+refines some subgoal using \hbox{\tt resolve_tac [{\it thm}]}
+\index{*resolve_tac}
+
+\item[\ttindexbold{fe} {\it thm};] 
+refines some subgoal using \hbox{\tt eresolve_tac [{\it thm}]}
+\index{*eresolve_tac}
+
+\item[\ttindexbold{fd} {\it thm};] 
+refines some subgoal using \hbox{\tt dresolve_tac [{\it thm}]}
+\index{*dresolve_tac}
+\end{description}
+
+\subsubsection{Resolution with a list of theorems}
+\begin{ttbox} 
+frs: thm list -> unit
+fes: thm list -> unit
+fds: thm list -> unit
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{frs} {\it thms};] 
+refines some subgoal using \hbox{\tt resolve_tac {\it thms}}
+\index{*resolve_tac}
+
+\item[\ttindexbold{fes} {\it thms};] 
+refines some subgoal using \hbox{\tt eresolve_tac {\it thms}} 
+\index{*eresolve_tac}
+
+\item[\ttindexbold{fds} {\it thms};] 
+refines some subgoal using \hbox{\tt dresolve_tac {\it thms}} 
+\index{*dresolve_tac}
+\end{description}
+
+\subsection{Other shortcuts}
+\begin{ttbox} 
+bw  : thm -> unit
+bws : thm list -> unit
+ren : string -> int -> unit
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{bw} {\it def};] 
+performs \hbox{\tt by (rewrite_goals_tac [{\it def}]);}
+It unfolds definitions in the subgoals (but not the main goal), by
+meta-rewriting with the given definition.\index{*rewrite_goals_tac}
+\index{rewriting!meta-level}
+
+\item[\ttindexbold{bws}] 
+is like {\tt bw} but takes a list of definitions.
+
+\item[\ttindexbold{ren} {\it names} {\it i};] 
+performs \hbox{\tt by (rename_tac {\it names} {\it i});} it renames
+parameters in subgoal~$i$.  (Ignore the message {\tt Warning:\ same as
+previous level}.)
+\index{*rename_tac}\index{parameters!renaming}
+\end{description}
+
+
+
+\section{Advanced features}
+\subsection{Executing batch proofs}
+\index{proofs!batch|bold}
+\index{batch execution|bold}
+\begin{ttbox} 
+prove_goal  :  theory->          string->(thm list->tactic list)->thm
+prove_goalw :  theory->thm list->string->(thm list->tactic list)->thm
+prove_goalw_cterm: thm list->Sign.cterm->(thm list->tactic list)->thm
+\end{ttbox}
+A collection of derivations may be stored for batch execution using these
+functions.  They create an initial proof state, then apply a tactic to it,
+yielding a sequence of final proof states.  The head of the sequence is
+returned, provided it is an instance of the theorem originally proposed.
+The forms {\tt prove_goal}, {\tt prove_goalw} and {\tt prove_goalw_cterm}
+are analogous to \ttindex{goal}, \ttindex{goalw} and \ttindex{goalw_cterm}.
+
+The tactic is specified by a function from theorem lists to tactic lists.
+The function is applied to the list of meta-hypotheses taken from the main
+goal.  The resulting tactics are applied in sequence (using {\tt EVERY}).
+For example, a proof consisting of the commands
+\begin{ttbox} 
+val prems = goal {\it theory} {\it formula};
+by \(tac@1\);  \ldots  by \(tac@n\);
+val my_thm = result();
+\end{ttbox}
+can be transformed to an expression as follows:
+\begin{ttbox} 
+val my_thm = prove_goal {\it theory} {\it formula}
+ (fn prems=> [ \(tac@1\), \ldots, \(tac@n\) ]);
+\end{ttbox}
+The methods perform identical processing of the initial {\it formula} and
+the final proof state, but {\tt prove_goal} executes the tactic as a
+atomic operation, bypassing the subgoal module.  The commands can also be
+encapsulated in an expression using~{\tt let}:
+\begin{ttbox} 
+val my_thm = 
+    let val prems = goal {\it theory} {\it formula}
+    in  by \(tac@1\);  \ldots  by \(tac@n\);  result() end;
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{prove_goal} {\it theory} {\it formula} {\it tacsf}; ] 
+executes a proof of the {\it formula\/} in the given {\it theory}, using
+the given tactic function.
+
+\item[\ttindexbold{prove_goalw} {\it theory} {\it defs} {\it formula} 
+      {\it tacsf}; ]
+is like {\tt prove_goal} but also applies the list of {\it defs\/} as
+meta-rewrite rules to the first subgoal and the premises.
+\index{rewriting!meta-level}
+
+\item[\ttindexbold{prove_goalw_cterm} {\it theory} {\it defs} {\it ct}
+      {\it tacsf}; ] 
+is a version of {\tt prove_goalw} for special applications.  The main
+goal is supplied as a cterm, not as a string.  Typically, the cterm is
+created from a term~$t$ by \hbox{\tt Sign.cterm_of (sign_of thy) $t$}.
+\index{*Sign.cterm_of}\index{*sign_of}
+\end{description}
+
+
+\subsection{Managing multiple proofs}
+\index{proofs!managing multiple|bold}
+You may save the current state of the subgoal module and resume work on it
+later.  This serves two purposes.  
+\begin{enumerate}
+\item At some point, you may be uncertain of the next step, and
+wish to experiment.
+
+\item During a proof, you may see that a lemma should be proved first.
+\end{enumerate}
+Each saved proof state consists of a list of levels; \ttindex{chop} behaves
+independently for each of the saved proofs.  In addition, each saved state
+carries a separate \ttindex{undo} list.
+
+\subsubsection{The stack of proof states}
+\index{proofs!stacking|bold}
+\begin{ttbox} 
+push_proof   : unit -> unit
+pop_proof    : unit -> thm list
+rotate_proof : unit -> thm list
+\end{ttbox}
+The subgoal module maintains a stack of proof states.  Most subgoal
+commands affect only the top of the stack.  The \ttindex{goal} command {\bf
+replaces} the top of the stack; the only command that pushes a proof on the
+stack is {\tt push_proof}.
+
+To save some point of the proof, call {\tt push_proof}.  You may now
+state a lemma using \ttindex{goal}, or simply continue to apply tactics.
+Later, you can return to the saved point by calling {\tt pop_proof} or 
+{\tt rotate_proof}. 
+
+To view the entire stack, call {\tt rotate_proof} repeatedly; as it rotates
+the stack, it prints the new top element.
+
+\begin{description}
+\item[\ttindexbold{push_proof}();]  
+duplicates the top element of the stack, pushing a copy of the current
+proof state on to the stack.
+
+\item[\ttindexbold{pop_proof}();]  
+discards the top element of the stack.  It returns the list of
+meta-hypotheses associated with the new proof;  you should bind these to an
+\ML\ identifier.  They can also be obtained by calling \ttindex{premises}.
+
+\item[\ttindexbold{rotate_proof}]  
+rotates the stack, moving the top element to the bottom.  It returns the
+list of assumptions associated with the new proof.
+\end{description}
+
+
+\subsubsection{Saving and restoring proof states}
+\index{proofs!saving and restoring|bold}
+\begin{ttbox} 
+save_proof    : unit -> proof
+restore_proof : proof -> thm list
+\end{ttbox}
+States of the subgoal module may be saved as \ML\ values of
+type~\ttindex{proof}, and later restored.
+
+\begin{description}
+\item[\ttindexbold{save_proof}();]  
+returns the current state, which is on top of the stack.
+
+\item[\ttindexbold{restore_proof} {\it prf}]  
+replaces the top of the stack by~{\it prf}.  It returns the list of
+assumptions associated with the new proof.
+\end{description}
+
+
+\subsection{Debugging and inspecting}
+\index{proofs!debugging|bold}
+\index{debugging}
+These specialized operations support the debugging of tactics.  They refer
+to the current proof state of the subgoal module, and fail if there is no
+proof underway.
+
+\subsubsection{Reading and printing terms}
+\index{reading!terms|bold}\index{printing!terms}\index{printing!types}
+\begin{ttbox} 
+read    : string -> term
+prin    : term -> unit
+printyp : typ -> unit
+\end{ttbox}
+These read and print terms (or types) using the syntax associated with the
+proof state.
+
+\begin{description}
+\item[\ttindexbold{read} {\it string}]  
+reads the {\it string} as a term, without type checking.
+
+\item[\ttindexbold{prin} {\it t};]  
+prints the term~$t$ at the terminal.
+
+\item[\ttindexbold{printyp} {\it T};]  
+prints the type~$T$ at the terminal.
+\end{description}
+
+\subsubsection{Inspecting the proof state}
+\index{proofs!inspecting the state|bold}
+\begin{ttbox} 
+topthm  : unit -> thm
+getgoal : int -> term
+gethyps : int -> thm list
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{topthm}()]  
+returns the proof state as an Isabelle theorem.  This is what \ttindex{by}
+would supply to a tactic at this point.  It omits the post-processing of
+\ttindex{result} and \ttindex{uresult}.
+
+\item[\ttindexbold{getgoal} {\it i}]  
+returns subgoal~$i$ of the proof state, as a term.  You may print
+this using {\tt prin}, though you may have to examine the internal
+data structure in order to locate the problem!
+
+\item[\ttindexbold{gethyps} {\it i}]  
+returns the hypotheses of subgoal~$i$ as meta-level assumptions.  In these
+theorems, the subgoal's parameters become free variables.  This command is
+supplied for debugging uses of \ttindex{METAHYPS}.
+\end{description}
+
+\subsubsection{Filtering lists of rules}
+\begin{ttbox} 
+filter_goal: (term*term->bool) -> thm list -> int -> thm list
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{filter_goal} {\it could} {\it ths} {\it i}] 
+applies \hbox{\tt filter_thms {\it could}} to subgoal~$i$ of the proof
+state and returns the list of theorems that survive the filtering. 
+\end{description}
+
+\index{subgoal module|)}
+\index{proofs|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/introduction.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,268 @@
+%% $Id$
+\maketitle 
+\pagenumbering{roman} \tableofcontents \clearfirst
+
+\chapter{Introduction} 
+\index{sessions|(}
+This manual is a comprehensive description of Isabelle, including all
+commands, functions and packages.  Please do not read it: it is intended
+for reference.  It is not a tutorial!  The manual assumes
+familiarity with the basic concepts explained in {\em Introduction to
+Isabelle}.
+
+Functions are organized by their purpose, by their operands (subgoals,
+tactics, theorems), and by their usefulness.  In each section, basic
+functions appear first, then advanced functions, and finally esoteric
+functions.  When you are looking for a way of performing some task, scan
+the Table of Contents for a relevant heading.
+
+The Index provides an alphabetical listing.  Page numbers of definitions
+are printed in {\bf bold}, passing references in Roman type.  Use the Index
+when you are looking for the definition of a particular Isabelle function.
+
+This manual contains few examples.  Many files of examples are distributed
+with Isabelle, however; please experiment interactively.
+
+
+\section{Basic interaction with Isabelle}
+\index{sessions!saving|bold}\index{saving your work|bold}
+Isabelle provides no means of storing theorems or proofs on files.
+Theorems are simply part of the \ML{} state and are named by \ML{}
+identifiers.  To save your work between sessions, you must save a copy of
+the \ML{} image.  The procedure for doing so is compiler-dependent:
+\begin{itemize}
+\item At the end of a session, Poly/\ML{} saves the state, provided you have
+created a database for your own use.  You can create a database by copying
+an existing one, or by calling the Poly/\ML{} function {\tt make_database};
+the latter course uses much less disc space.  Note that a Poly/\ML{}
+database {\bf does not} save the contents of references, such as the
+current state of a backward proof.
+
+\item With New Jersey \ML{} you must save the state explicitly before
+ending the session.  While Poly/\ML{} database can be small, a New Jersey
+image occupies several megabytes.
+\end{itemize}
+See your \ML{} compiler's documentation for full instructions on saving the
+state.
+
+Saving the state is not enough.  Record, on a file, the top-level commands
+that generate your theories and proofs.  Such a record allows you to replay
+the proofs whenever required, for instance after making minor changes to
+the axioms.  Ideally, your record will be intelligible to others as a
+formal description of your work.
+
+Since Isabelle's user interface is the \ML{} top level, some kind of window
+support is essential.  One window displays the Isabelle session, while the
+other displays a file --- your proof record --- being edited.  Some people
+use a screen editor such as Emacs, which supports windows and can manage
+interactive sessions.  Others prefer to use a workstation running the X Window
+System.
+
+
+\section{Ending a session}
+\index{sessions!ending|bold}
+\begin{ttbox} 
+quit     : unit -> unit
+commit   : unit -> unit \hfill{\bf Poly/ML only}
+exportML : string -> bool \hfill{\bf New Jersey ML only}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{quit}();]  
+aborts the Isabelle session, without saving the state.
+
+\item[\ttindexbold{commit}();]  saves the current state in your
+Poly/\ML{} database without finishing the session.  Values of reference
+variables are lost, so never do this during an interactive proof!
+
+\item[\ttindexbold{exportML} \tt"{\it file}";]  saves an
+image of your session to the given {\it file}.
+\end{description}
+
+\begin{warn}
+Typing control-D also finishes the session, but its effect is
+compiler-dependent.  Poly/\ML{} will then save the state, if you have a
+private database.  New Jersey \ML{} will discard the state!
+\end{warn}
+
+
+\section{Reading files of proofs and theories}
+\index{files, reading|bold}
+\begin{ttbox} 
+cd           : string -> unit
+use          : string -> unit
+use_thy      : string -> unit
+time_use     : string -> unit
+time_use_thy : string -> unit
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{cd} \tt"{\it dir}";]  changes to {\it dir\/} the default
+directory for reading files.
+
+\item[\ttindexbold{use} \tt"$file$";]  
+reads the given {\it file} as input to the \ML{} session.  Reading a file
+of Isabelle commands is the usual way of replaying a proof.
+
+\item[\ttindexbold{use_thy} \tt"$tname$";]  
+reads a theory definition from file {\it tname}{\tt.thy} and also reads
+{\ML} commands from the file {\it tname}{\tt.ML}, if it exists.  See
+Chapter~\ref{theories} for details.
+
+\item[\ttindexbold{time_use} \tt"$file$";]  
+performs {\tt use~"$file$"} and prints the total execution time.
+
+\item[\ttindexbold{time_use_thy} \tt"$tname$";]  
+performs {\tt use_thy "$tname$"} and prints the total execution time.
+\end{description}
+
+
+\section{Printing of terms and theorems}
+\index{printing!flags|bold}
+Isabelle's pretty printer is controlled by a number of parameters.
+
+\subsection{Printing limits}
+\begin{ttbox} 
+Pretty.setdepth  : int -> unit
+Pretty.setmargin : int -> unit
+print_depth      : int -> unit
+\end{ttbox}
+These set limits for terminal output.
+
+\begin{description}
+\item[\ttindexbold{Pretty.setdepth} \(d\);]  tells
+Isabelle's pretty printer to limit the printing depth to~$d$.  This affects
+Isabelle's display of theorems and terms.  The default value is~0, which
+permits printing to an arbitrary depth.  Useful values for $d$ are~10 and~20.
+
+\item[\ttindexbold{Pretty.setmargin} \(m\);]  tells
+Isabelle's pretty printer to assume a right margin (page width) of~$m$.
+The initial margin is~80.
+
+\item[\ttindexbold{print_depth} \(n\);]  limits
+the printing depth of complex \ML{} values, such as theorems and terms.
+This command affects the \ML{} top level and its effect is
+compiler-dependent.  Typically $n$ should be less than~10.
+\end{description}
+
+
+\subsection{Printing of meta-level hypotheses}
+\index{hypotheses!meta-level!printing of|bold}
+\begin{ttbox} 
+show_hyps: bool ref \hfill{\bf initially true}
+\end{ttbox}
+A theorem's hypotheses are normally displayed, since such dependence is
+important.  If this information becomes too verbose, however, it can be
+switched off;  each hypothesis is then displayed as a dot.
+\begin{description}
+\item[\ttindexbold{show_hyps} \tt:= true;]   
+makes Isabelle show meta-level hypotheses when printing a theorem; setting
+it to {\tt false} suppresses them.
+\end{description}
+
+
+\subsection{Printing of types and sorts}
+\begin{ttbox} 
+show_types: bool ref \hfill{\bf initially false}
+show_sorts: bool ref \hfill{\bf initially false}
+\end{ttbox}
+These control Isabelle's display of types and sorts.  Normally terms are
+printed without type and sorts because they are verbose.  Occasionally you
+may require this information, say to discover why a polymorphic inference rule
+fails to resolve with some goal.
+
+\begin{description}
+\item[\ttindexbold{show_types} \tt:= true;]\index{types}
+makes Isabelle show types when printing a term or theorem.
+
+\item[\ttindexbold{show_sorts} \tt:= true;]\index{sorts}
+makes Isabelle show the sorts of type variables.  It has no effect unless
+{\tt show_types} is~{\tt true}. 
+\end{description}
+
+
+\subsection{$\eta$-contraction before printing}
+\begin{ttbox} 
+eta_contract: bool ref \hfill{\bf initially false}
+\end{ttbox}
+The {\bf $\eta$-contraction law} asserts $(\lambda x.f(x))\equiv f$,
+provided $x$ is not free in ~$f$.  It asserts {\bf extensionality} of
+functions: $f\equiv g$ if $f(x)\equiv g(x)$ for all~$x$.  Higher-order
+unification puts terms into a fully $\eta$-expanded form.  For example, if
+$F$ has type $(\tau\To\tau)\To\tau$ then its expanded form is $\lambda
+h.F(\lambda x.h(x))$.  By default, the user sees this expanded form.
+
+\begin{description}
+\item[\ttindexbold{eta_contract} \tt:= true;]
+makes Isabelle perform $\eta$-contractions before printing, so that
+$\lambda h.F(\lambda x.h(x))$ appears simply as~$F$.  The
+distinction between a term and its $\eta$-expanded form occasionally
+matters.
+\end{description}
+
+
+\section{Displaying exceptions as error messages}
+\index{printing!exceptions|bold}\index{exceptions|bold}
+\begin{ttbox} 
+print_exn: exn -> 'a
+\end{ttbox}
+Certain Isabelle primitives, such as the forward proof functions {\tt RS}
+and {\tt RSN}, are called both interactively and from programs.  They
+indicate errors not by printing messages, but by raising exceptions.  For
+interactive use, \ML's reporting of an uncaught exception is most
+uninformative.
+
+\begin{description}
+\item[\ttindexbold{print_exn} $e$] 
+displays the exception~$e$ in a readable manner, and then re-raises~$e$.
+Typical usage is~\hbox{\tt \ldots{} handle e => print_exn e;}, where
+\ldots{} is an expression that may raise an exception.
+
+{\tt print_exn} can display the following common exceptions, which concern
+types, terms, theorems and theories, respectively.  Each carries a message
+and related information.
+\begin{ttbox} 
+exception TYPE   of string * typ list * term list
+exception TERM   of string * term list
+exception THM    of string * int * thm list
+exception THEORY of string * theory list
+\end{ttbox}
+{\tt print_exn} calls \ttindex{prin} to print terms.  This obtains pretty
+printing information from the proof state last stored in the subgoal
+module, and will fail if no interactive proof has begun in the current
+session.
+\end{description}
+
+
+\section{Shell scripts}
+\index{shell scripts|bold}
+The following files are distributed with Isabelle, and work under
+Unix$^{\rm TM}$.  They can be executed as commands to the Unix shell.  Some
+of them depend upon shell environment variables.
+\begin{description}
+\item[\ttindexbold{make-all} $switches$] 
+compiles the Isabelle system, when executed on the source directory.
+Environment variables specify which \ML{} compiler (and {\tt Makefile}s) to
+use.  These variables, and the {\it switches}, are documented on the file
+itself.
+
+\item[\ttindexbold{teeinput} $program$] 
+executes the {\it program}, while piping the standard input to a log file
+designated by the \verb|$LISTEN| environment variable.  Normally the
+program is Isabelle, and the log file receives a copy of all the Isabelle
+commands.
+
+\item[\ttindexbold{xlisten} $program$] 
+  is a trivial `user interface' for the X Window System.  It creates two
+  windows using {\tt xterm}.  One executes an interactive session via
+  \hbox{\tt teeinput $program$}, while the other displays the log file.  To
+  make a proof record, simply paste lines from the log file into an editor
+  window.
+
+\item[\ttindexbold{expandshort} $files$] 
+  reads the {\it files\/} and replaces all occurrences of the shorthand
+  commands {\tt br}, {\tt be}, {\tt brs}, {\tt bes}, etc., with commands
+  like \hbox{\tt by (resolve_tac \ldots)}.  The commands should appear one
+  per line.  The old versions of the files
+  are renamed to have the suffix~\verb'~~'.
+\end{description}
+
+\index{sessions|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/ref.bbl	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,25 @@
+\begin{thebibliography}{1}
+
+\bibitem{charniak80}
+E.~Charniak, C.~K. Riesbeck, and D.~V. McDermott.
+\newblock {\em Artificial Intelligence Programming}.
+\newblock Lawrence Erlbaum Associates, 1980.
+
+\bibitem{debruijn72}
+N.~G. de~Bruijn.
+\newblock Lambda calculus notation with nameless dummies, a tool for automatic
+  formula manipulation, with application to the {Church-Rosser Theorem}.
+\newblock {\em Indagationes Mathematicae}, 34:381--392, 1972.
+
+\bibitem{paulson91}
+Lawrence~C. Paulson.
+\newblock {\em {ML} for the Working Programmer}.
+\newblock Cambridge University Press, 1991.
+
+\bibitem{pelletier86}
+F.~J. Pelletier.
+\newblock Seventy-five problems for testing automatic theorem provers.
+\newblock {\em Journal of Automated Reasoning}, 2:191--216, 1986.
+\newblock Errata, JAR 4 (1988), 235--236.
+
+\end{thebibliography}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/ref.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,51 @@
+\documentstyle[a4,12pt,proof,iman,alltt]{report}
+%% $Id$
+%%% \includeonly{thm}
+%%% to index ids: \[\\tt \([a-zA-Z0-9][a-zA-Z0-9_'.]*\)    [\\ttindexbold{\1}
+%%% to delete old ones:  \\indexbold{\*[^}]*}
+%% run    sedindex ref    to prepare index file
+%%% needs chapter on Provers/typedsimp.ML?
+\title{The Isabelle Reference Manual}
+
+\author{{\em Lawrence C. Paulson}\thanks
+{Tobias Nipkow, of T. U. Munich, wrote Chapter~8 and part of Chapter~6.
+ Sara Kalvala and Markus Wenzel suggested changes and corrections.
+ The research has been funded by the SERC (grants GR/G53279, GR/H40570)
+  and by ESPRIT project 6453: Types).} 
+\\  
+        Computer Laboratory \\ University of Cambridge \\[2ex]
+        {\small{\em Electronic mail\/}: {\tt lcp@cl.cam.ac.uk}} \\[3cm]
+    {\small Copyright \copyright{} \number\year{} by Lawrence C. Paulson}
+}
+\date{} 
+\makeindex
+
+\underscoreoff
+
+\setcounter{secnumdepth}{1} \setcounter{tocdepth}{2}
+
+\pagestyle{headings}
+\sloppy
+\binperiod     %%%treat . like a binary operator
+
+\begin{document}
+\index{definitions|see{rewriting, meta-level}}
+\index{rewriting!object-level|see{simplification}}
+\index{rules!meta-level|see{meta-rules}}
+\include{introduction}
+\include{goals}
+\include{tactic}
+\include{tctical}
+\include{thm}
+\include{theories}
+\include{substitution}
+\include{simplifier}
+\include{classical}
+
+%%seealso's must be last so that they appear last in the index entries
+\index{rewriting!meta-level|seealso{tactics, theorems}}
+
+\bibliographystyle{plain} \small\raggedright\frenchspacing
+\bibliography{atp,funprog,general,logicprog,theory}
+\input{ref.ind}
+\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/ref.toc	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,143 @@
+\contentsline {chapter}{\numberline {1}Introduction}{1}
+\contentsline {section}{\numberline {1.1}Basic interaction with Isabelle}{1}
+\contentsline {section}{\numberline {1.2}Ending a session}{2}
+\contentsline {section}{\numberline {1.3}Reading files of proofs and theories}{2}
+\contentsline {section}{\numberline {1.4}Printing of terms and theorems}{2}
+\contentsline {subsection}{Printing limits}{3}
+\contentsline {subsection}{Printing of meta-level hypotheses}{3}
+\contentsline {subsection}{Printing of types and sorts}{3}
+\contentsline {subsection}{$\eta $-contraction before printing}{3}
+\contentsline {section}{\numberline {1.5}Displaying exceptions as error messages}{4}
+\contentsline {section}{\numberline {1.6}Shell scripts}{4}
+\contentsline {chapter}{\numberline {2}Proof Management: The Subgoal Module}{6}
+\contentsline {section}{\numberline {2.1}Basic commands}{6}
+\contentsline {subsection}{Starting a backward proof}{6}
+\contentsline {subsection}{Applying a tactic}{7}
+\contentsline {subsection}{Extracting the proved theorem}{8}
+\contentsline {subsection}{Undoing and backtracking}{8}
+\contentsline {subsection}{Printing the proof state}{9}
+\contentsline {subsection}{Timing}{9}
+\contentsline {section}{\numberline {2.2}Shortcuts for applying tactics}{9}
+\contentsline {subsection}{Refining a given subgoal}{9}
+\contentsline {subsubsection}{Resolution with a list of theorems}{9}
+\contentsline {subsection}{Scanning shortcuts}{10}
+\contentsline {subsubsection}{Proof by assumption and resolution}{10}
+\contentsline {subsubsection}{Resolution with a list of theorems}{10}
+\contentsline {subsection}{Other shortcuts}{10}
+\contentsline {section}{\numberline {2.3}Advanced features}{11}
+\contentsline {subsection}{Executing batch proofs}{11}
+\contentsline {subsection}{Managing multiple proofs}{11}
+\contentsline {subsubsection}{The stack of proof states}{12}
+\contentsline {subsubsection}{Saving and restoring proof states}{12}
+\contentsline {subsection}{Debugging and inspecting}{12}
+\contentsline {subsubsection}{Reading and printing terms}{13}
+\contentsline {subsubsection}{Inspecting the proof state}{13}
+\contentsline {subsubsection}{Filtering lists of rules}{13}
+\contentsline {chapter}{\numberline {3}Tactics}{14}
+\contentsline {section}{\numberline {3.1}Resolution and assumption tactics}{14}
+\contentsline {subsection}{Resolution tactics}{14}
+\contentsline {subsection}{Assumption tactics}{15}
+\contentsline {subsection}{Matching tactics}{15}
+\contentsline {subsection}{Resolution with instantiation}{15}
+\contentsline {section}{\numberline {3.2}Other basic tactics}{16}
+\contentsline {subsection}{Definitions and meta-level rewriting}{16}
+\contentsline {subsection}{Tactic shortcuts}{17}
+\contentsline {subsection}{Inserting premises and facts}{17}
+\contentsline {subsection}{Theorems useful with tactics}{18}
+\contentsline {section}{\numberline {3.3}Obscure tactics}{18}
+\contentsline {subsection}{Tidying the proof state}{18}
+\contentsline {subsection}{Renaming parameters in a goal}{18}
+\contentsline {subsection}{Composition: resolution without lifting}{19}
+\contentsline {section}{\numberline {3.4}Managing lots of rules}{19}
+\contentsline {subsection}{Combined resolution and elim-resolution}{20}
+\contentsline {subsection}{Discrimination nets for fast resolution}{20}
+\contentsline {section}{\numberline {3.5}Programming tools for proof strategies}{21}
+\contentsline {subsection}{Operations on type {\ptt tactic}}{22}
+\contentsline {subsection}{Tracing}{22}
+\contentsline {subsection}{Sequences}{23}
+\contentsline {subsubsection}{Basic operations on sequences}{23}
+\contentsline {subsubsection}{Converting between sequences and lists}{23}
+\contentsline {subsubsection}{Combining sequences}{23}
+\contentsline {chapter}{\numberline {4}Tacticals}{25}
+\contentsline {section}{\numberline {4.1}The basic tacticals}{25}
+\contentsline {subsection}{Joining two tactics}{25}
+\contentsline {subsection}{Joining a list of tactics}{25}
+\contentsline {subsection}{Repetition tacticals}{26}
+\contentsline {subsection}{Identities for tacticals}{26}
+\contentsline {section}{\numberline {4.2}Control and search tacticals}{27}
+\contentsline {subsection}{Filtering a tactic's results}{27}
+\contentsline {subsection}{Depth-first search}{28}
+\contentsline {subsection}{Other search strategies}{28}
+\contentsline {subsection}{Auxiliary tacticals for searching}{29}
+\contentsline {subsection}{Predicates and functions useful for searching}{29}
+\contentsline {section}{\numberline {4.3}Tacticals for subgoal numbering}{29}
+\contentsline {subsection}{Restricting a tactic to one subgoal}{30}
+\contentsline {subsection}{Scanning for a subgoal by number}{31}
+\contentsline {subsection}{Joining tactic functions}{32}
+\contentsline {subsection}{Applying a list of tactics to 1}{32}
+\contentsline {chapter}{\numberline {5}Theorems and Forward Proof}{33}
+\contentsline {section}{\numberline {5.1}Basic operations on theorems}{33}
+\contentsline {subsection}{Pretty-printing a theorem}{33}
+\contentsline {subsubsection}{Top-level commands}{33}
+\contentsline {subsubsection}{Operations for programming}{34}
+\contentsline {subsection}{Forwards proof: joining rules by resolution}{34}
+\contentsline {subsection}{Expanding definitions in theorems}{35}
+\contentsline {subsection}{Instantiating a theorem}{35}
+\contentsline {subsection}{Miscellaneous forward rules}{35}
+\contentsline {subsection}{Taking a theorem apart}{36}
+\contentsline {subsection}{Tracing flags for unification}{36}
+\contentsline {section}{\numberline {5.2}Primitive meta-level inference rules}{37}
+\contentsline {subsection}{Propositional rules}{38}
+\contentsline {subsubsection}{Assumption}{38}
+\contentsline {subsubsection}{Implication}{38}
+\contentsline {subsubsection}{Logical equivalence (equality)}{39}
+\contentsline {subsection}{Equality rules}{39}
+\contentsline {subsection}{The $\lambda $-conversion rules}{39}
+\contentsline {subsection}{Universal quantifier rules}{40}
+\contentsline {subsubsection}{Forall introduction}{40}
+\contentsline {subsubsection}{Forall elimination}{40}
+\contentsline {subsubsection}{Instantiation of unknowns}{41}
+\contentsline {subsection}{Freezing/thawing type variables}{41}
+\contentsline {section}{\numberline {5.3}Derived rules for goal-directed proof}{41}
+\contentsline {subsection}{Proof by assumption}{41}
+\contentsline {subsection}{Resolution}{41}
+\contentsline {subsection}{Composition: resolution without lifting}{42}
+\contentsline {subsection}{Other meta-rules}{42}
+\contentsline {chapter}{\numberline {6}Theories, Terms and Types}{44}
+\contentsline {section}{\numberline {6.1}Defining Theories}{44}
+\contentsline {section}{\numberline {6.2}Basic operations on theories}{47}
+\contentsline {subsection}{Extracting an axiom from a theory}{47}
+\contentsline {subsection}{Building a theory}{47}
+\contentsline {subsection}{Inspecting a theory}{48}
+\contentsline {section}{\numberline {6.3}Terms}{49}
+\contentsline {section}{\numberline {6.4}Certified terms}{50}
+\contentsline {subsection}{Printing terms}{50}
+\contentsline {subsection}{Making and inspecting certified terms}{50}
+\contentsline {section}{\numberline {6.5}Types}{51}
+\contentsline {section}{\numberline {6.6}Certified types}{51}
+\contentsline {subsection}{Printing types}{51}
+\contentsline {subsection}{Making and inspecting certified types}{51}
+\contentsline {chapter}{\numberline {7}Substitution Tactics}{53}
+\contentsline {section}{\numberline {7.1}Simple substitution}{53}
+\contentsline {section}{\numberline {7.2}Substitution in the hypotheses}{54}
+\contentsline {section}{\numberline {7.3}Setting up {\ptt hyp_subst_tac}}{54}
+\contentsline {chapter}{\numberline {8}Simplification}{57}
+\contentsline {section}{\numberline {8.1}Simplification sets}{57}
+\contentsline {subsection}{Rewrite rules}{57}
+\contentsline {subsection}{Congruence rules}{58}
+\contentsline {subsection}{The subgoaler}{58}
+\contentsline {subsection}{The solver}{59}
+\contentsline {subsection}{The looper}{59}
+\contentsline {section}{\numberline {8.2}The simplification tactics}{59}
+\contentsline {section}{\numberline {8.3}Example: using the simplifier}{60}
+\contentsline {chapter}{\numberline {9}The classical theorem prover}{64}
+\contentsline {section}{\numberline {9.1}The sequent calculus}{64}
+\contentsline {section}{\numberline {9.2}Simulating sequents by natural deduction}{65}
+\contentsline {section}{\numberline {9.3}Extra rules for the sequent calculus}{66}
+\contentsline {section}{\numberline {9.4}Classical rule sets}{67}
+\contentsline {section}{\numberline {9.5}The classical tactics}{68}
+\contentsline {subsection}{Single-step tactics}{69}
+\contentsline {subsection}{The automatic tactics}{69}
+\contentsline {subsection}{Other useful tactics}{69}
+\contentsline {subsection}{Creating swapped rules}{70}
+\contentsline {section}{\numberline {9.6}Setting up the classical prover}{70}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/simp.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,512 @@
+%%%THIS DOCUMENTS THE OBSOLETE SIMPLIFIER!!!!
+\chapter{Simplification} \label{simp-chap}
+\index{simplification|(}
+Object-level rewriting is not primitive in Isabelle.  For efficiency,
+perhaps it ought to be.  On the other hand, it is difficult to conceive of
+a general mechanism that could accommodate the diversity of rewriting found
+in different logics.  Hence rewriting in Isabelle works via resolution,
+using unknowns as place-holders for simplified terms.  This chapter
+describes a generic simplification package, the functor~\ttindex{SimpFun},
+which expects the basic laws of equational logic and returns a suite of
+simplification tactics.  The code lives in
+\verb$Provers/simp.ML$.
+
+This rewriting package is not as general as one might hope (using it for {\tt
+HOL} is not quite as convenient as it could be; rewriting modulo equations is
+not supported~\ldots) but works well for many logics.  It performs
+conditional and unconditional rewriting and handles multiple reduction
+relations and local assumptions.  It also has a facility for automatic case
+splits by expanding conditionals like {\it if-then-else\/} during rewriting.
+
+For many of Isabelle's logics ({\tt FOL}, {\tt ZF}, {\tt LCF} and {\tt HOL})
+the simplifier has been set up already. Hence we start by describing the
+functions provided by the simplifier --- those functions exported by
+\ttindex{SimpFun} through its result signature \ttindex{SIMP} shown in
+Figure~\ref{SIMP}.  
+
+
+\section{Simplification sets}
+\index{simplification sets}
+The simplification tactics are controlled by {\bf simpsets}, which consist of
+three things:
+\begin{enumerate}
+\item {\bf Rewrite rules}, which are theorems like 
+$\Var{m} + succ(\Var{n}) = succ(\Var{m} + \Var{n})$.  {\bf Conditional}
+rewrites such as $m<n \Imp m/n = 0$ are permitted.
+\index{rewrite rules}
+
+\item {\bf Congruence rules}, which typically have the form
+\index{congruence rules}
+\[ \List{\Var{x@1} = \Var{y@1}; \ldots; \Var{x@n} = \Var{y@n}} \Imp
+   f(\Var{x@1},\ldots,\Var{x@n}) = f(\Var{y@1},\ldots,\Var{y@n}).
+\]
+
+\item The {\bf auto-tactic}, which attempts to solve the simplified
+subgoal, say by recognizing it as a tautology.
+\end{enumerate}
+
+\subsection{Congruence rules}
+Congruence rules enable the rewriter to simplify subterms.  Without a
+congruence rule for the function~$g$, no argument of~$g$ can be rewritten.
+Congruence rules can be generalized in the following ways:
+
+{\bf Additional assumptions} are allowed:
+\[ \List{\Var{P@1} \bimp \Var{Q@1};\; \Var{Q@1} \Imp \Var{P@2} \bimp \Var{Q@2}}
+   \Imp (\Var{P@1} \imp \Var{P@2}) \bimp (\Var{Q@1} \imp \Var{Q@2})
+\]
+This rule assumes $Q@1$, and any rewrite rules it contains, while
+simplifying~$P@2$.  Such ``local'' assumptions are effective for rewriting
+formulae such as $x=0\imp y+x=y$.
+
+{\bf Additional quantifiers} are allowed, typically for binding operators:
+\[ \List{\Forall z. \Var{P}(z) \bimp \Var{Q}(z)} \Imp
+   \forall x.\Var{P}(x) \bimp \forall x.\Var{Q}(x)
+\]
+
+{\bf Different equalities} can be mixed.  The following example
+enables the transition from formula rewriting to term rewriting:
+\[ \List{\Var{x@1}=\Var{y@1};\Var{x@2}=\Var{y@2}} \Imp
+   (\Var{x@1}=\Var{x@2}) \bimp (\Var{y@1}=\Var{y@2})
+\]
+\begin{warn}
+It is not necessary to assert a separate congruence rule for each constant,
+provided your logic contains suitable substitution rules. The function {\tt
+mk_congs} derives congruence rules from substitution
+rules~\S\ref{simp-tactics}.
+\end{warn}
+
+
+\begin{figure}
+\indexbold{*SIMP}
+\begin{ttbox}
+infix 4 addrews addcongs delrews delcongs setauto;
+signature SIMP =
+sig
+  type simpset
+  val empty_ss  : simpset
+  val addcongs  : simpset * thm list -> simpset
+  val addrews   : simpset * thm list -> simpset
+  val delcongs  : simpset * thm list -> simpset
+  val delrews   : simpset * thm list -> simpset
+  val print_ss  : simpset -> unit
+  val setauto   : simpset * (int -> tactic) -> simpset
+  val ASM_SIMP_CASE_TAC : simpset -> int -> tactic
+  val ASM_SIMP_TAC      : simpset -> int -> tactic
+  val CASE_TAC          : simpset -> int -> tactic
+  val SIMP_CASE2_TAC    : simpset -> int -> tactic
+  val SIMP_THM          : simpset -> thm -> thm
+  val SIMP_TAC          : simpset -> int -> tactic
+  val SIMP_CASE_TAC     : simpset -> int -> tactic
+  val mk_congs          : theory -> string list -> thm list
+  val mk_typed_congs    : theory -> (string*string) list -> thm list
+  val tracing   : bool ref
+end;
+\end{ttbox}
+\caption{The signature {\tt SIMP}} \label{SIMP}
+\end{figure}
+
+
+\subsection{The abstract type {\tt simpset}}\label{simp-simpsets}
+Simpsets are values of the abstract type \ttindexbold{simpset}.  They are
+manipulated by the following functions:
+\index{simplification sets|bold}
+\begin{description}
+\item[\ttindexbold{empty_ss}] 
+is the empty simpset.  It has no congruence or rewrite rules and its
+auto-tactic always fails.
+
+\item[\tt $ss$ \ttindexbold{addcongs} $thms$] 
+is the simpset~$ss$ plus the congruence rules~$thms$.
+
+\item[\tt $ss$ \ttindexbold{delcongs} $thms$] 
+is the simpset~$ss$ minus the congruence rules~$thms$.
+
+\item[\tt $ss$ \ttindexbold{addrews} $thms$] 
+is the simpset~$ss$ plus the rewrite rules~$thms$.
+
+\item[\tt $ss$ \ttindexbold{delrews} $thms$] 
+is the simpset~$ss$ minus the rewrite rules~$thms$.
+
+\item[\tt $ss$ \ttindexbold{setauto} $tacf$] 
+is the simpset~$ss$ with $tacf$ for its auto-tactic.
+
+\item[\ttindexbold{print_ss} $ss$] 
+prints all the congruence and rewrite rules in the simpset~$ss$.
+\end{description}
+Adding a rule to a simpset already containing it, or deleting one
+from a simpset not containing it, generates a warning message.
+
+In principle, any theorem can be used as a rewrite rule.  Before adding a
+theorem to a simpset, {\tt addrews} preprocesses the theorem to extract the
+maximum amount of rewriting from it.  Thus it need not have the form $s=t$.
+In {\tt FOL} for example, an atomic formula $P$ is transformed into the
+rewrite rule $P \bimp True$.  This preprocessing is not fixed but logic
+dependent.  The existing logics like {\tt FOL} are fairly clever in this
+respect.  For a more precise description see {\tt mk_rew_rules} in
+\S\ref{SimpFun-input}.  
+
+The auto-tactic is applied after simplification to solve a goal.  This may
+be the overall goal or some subgoal that arose during conditional
+rewriting.  Calling ${\tt auto_tac}~i$ must either solve exactly
+subgoal~$i$ or fail.  If it succeeds without reducing the number of
+subgoals by one, havoc and strange exceptions may result.
+A typical auto-tactic is {\tt ares_tac [TrueI]}, which attempts proof by
+assumption and resolution with the theorem $True$.  In explicitly typed
+logics, the auto-tactic can be used to solve simple type checking
+obligations.  Some applications demand a sophisticated auto-tactic such as
+{\tt fast_tac}, but this could make simplification slow.
+  
+\begin{warn}
+Rewriting never instantiates unknowns in subgoals.  (It uses
+\ttindex{match_tac} rather than \ttindex{resolve_tac}.)  However, the
+auto-tactic is permitted to instantiate unknowns.
+\end{warn}
+
+
+\section{The simplification tactics} \label{simp-tactics}
+\index{simplification!tactics|bold}
+\index{tactics!simplification|bold}
+The actual simplification work is performed by the following tactics.  The
+rewriting strategy is strictly bottom up.  Conditions in conditional rewrite
+rules are solved recursively before the rewrite rule is applied.
+
+There are two basic simplification tactics:
+\begin{description}
+\item[\ttindexbold{SIMP_TAC} $ss$ $i$] 
+simplifies subgoal~$i$ using the rules in~$ss$.  It may solve the
+subgoal completely if it has become trivial, using the auto-tactic
+(\S\ref{simp-simpsets}).
+  
+\item[\ttindexbold{ASM_SIMP_TAC}] 
+is like \verb$SIMP_TAC$, but also uses assumptions as additional
+rewrite rules.
+\end{description}
+Many logics have conditional operators like {\it if-then-else}.  If the
+simplifier has been set up with such case splits (see~\ttindex{case_splits}
+in \S\ref{SimpFun-input}), there are tactics which automatically alternate
+between simplification and case splitting:
+\begin{description}
+\item[\ttindexbold{SIMP_CASE_TAC}] 
+is like {\tt SIMP_TAC} but also performs automatic case splits.
+More precisely, after each simplification phase the tactic tries to apply a
+theorem in \ttindex{case_splits}.  If this succeeds, the tactic calls
+itself recursively on the result.
+
+\item[\ttindexbold{ASM_SIMP_CASE_TAC}] 
+is like {\tt SIMP_CASE_TAC}, but also uses assumptions for
+rewriting.
+
+\item[\ttindexbold{SIMP_CASE2_TAC}] 
+is like {\tt SIMP_CASE_TAC}, but also tries to solve the
+pre-conditions of conditional simplification rules by repeated case splits.
+
+\item[\ttindexbold{CASE_TAC}] 
+tries to break up a goal using a rule in
+\ttindex{case_splits}.
+
+\item[\ttindexbold{SIMP_THM}] 
+simplifies a theorem using assumptions and case splitting.
+\end{description}
+Finally there are two useful functions for generating congruence
+rules for constants and free variables:
+\begin{description}
+\item[\ttindexbold{mk_congs} $thy$ $cs$] 
+computes a list of congruence rules, one for each constant in $cs$.
+Remember that the name of an infix constant
+\verb$+$ is \verb$op +$.
+
+\item[\ttindexbold{mk_typed_congs}] 
+computes congruence rules for explicitly typed free variables and
+constants.  Its second argument is a list of name and type pairs.  Names
+can be either free variables like {\tt P}, or constants like \verb$op =$.
+For example in {\tt FOL}, the pair
+\verb$("f","'a => 'a")$ generates the rule \verb$?x = ?y ==> f(?x) = f(?y)$.
+Such congruence rules are necessary for goals with free variables whose
+arguments need to be rewritten.
+\end{description}
+Both functions work correctly only if {\tt SimpFun} has been supplied with the
+necessary substitution rules.  The details are discussed in
+\S\ref{SimpFun-input} under {\tt subst_thms}.
+\begin{warn}
+Using the simplifier effectively may take a bit of experimentation. In
+particular it may often happen that simplification stops short of what you
+expected or runs forever. To diagnose these problems, the simplifier can be
+traced. The reference variable \ttindexbold{tracing} controls the output of
+tracing information.
+\index{tracing!of simplification}
+\end{warn}
+
+
+\section{Example: using the simplifier}
+\index{simplification!example}
+Assume we are working within {\tt FOL} and that
+\begin{description}
+\item[\tt Nat.thy] is a theory including the constants $0$, $Suc$ and $+$,
+\item[\tt add_0] is the rewrite rule $0+n = n$,
+\item[\tt add_Suc] is the rewrite rule $Suc(m)+n = Suc(m+n)$,
+\item[\tt induct] is the induction rule
+$\List{P(0); \Forall x. P(x)\Imp P(Suc(x))} \Imp P(n)$.
+\item[\tt FOL_ss] is a basic simpset for {\tt FOL}.
+\end{description}
+We generate congruence rules for $Suc$ and for the infix operator~$+$:
+\begin{ttbox}
+val nat_congs = mk_congs Nat.thy ["Suc", "op +"];
+prths nat_congs;
+{\out ?Xa = ?Ya ==> Suc(?Xa) = Suc(?Ya)}
+{\out [| ?Xa = ?Ya; ?Xb = ?Yb |] ==> ?Xa + ?Xb = ?Ya + ?Yb}
+\end{ttbox}
+We create a simpset for natural numbers by extending~{\tt FOL_ss}:
+\begin{ttbox}
+val add_ss = FOL_ss  addcongs nat_congs  
+                     addrews  [add_0, add_Suc];
+\end{ttbox}
+Proofs by induction typically involve simplification:\footnote
+{These examples reside on the file {\tt FOL/ex/nat.ML}.} 
+\begin{ttbox}
+goal Nat.thy "m+0 = m";
+{\out Level 0}
+{\out m + 0 = m}
+{\out  1. m + 0 = m}
+\ttbreak
+by (res_inst_tac [("n","m")] induct 1);
+{\out Level 1}
+{\out m + 0 = m}
+{\out  1. 0 + 0 = 0}
+{\out  2. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)}
+\end{ttbox}
+Simplification solves the first subgoal:
+\begin{ttbox}
+by (SIMP_TAC add_ss 1);
+{\out Level 2}
+{\out m + 0 = m}
+{\out  1. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)}
+\end{ttbox}
+The remaining subgoal requires \ttindex{ASM_SIMP_TAC} in order to use the
+induction hypothesis as a rewrite rule:
+\begin{ttbox}
+by (ASM_SIMP_TAC add_ss 1);
+{\out Level 3}
+{\out m + 0 = m}
+{\out No subgoals!}
+\end{ttbox}
+The next proof is similar.
+\begin{ttbox}
+goal Nat.thy "m+Suc(n) = Suc(m+n)";
+{\out Level 0}
+{\out m + Suc(n) = Suc(m + n)}
+{\out  1. m + Suc(n) = Suc(m + n)}
+\ttbreak
+by (res_inst_tac [("n","m")] induct 1);
+{\out Level 1}
+{\out m + Suc(n) = Suc(m + n)}
+{\out  1. 0 + Suc(n) = Suc(0 + n)}
+{\out  2. !!x. x + Suc(n) = Suc(x + n) ==> Suc(x) + Suc(n) = Suc(Suc(x) + n)}
+\end{ttbox}
+Using the tactical \ttindex{ALLGOALS}, a single command simplifies all the
+subgoals:
+\begin{ttbox}
+by (ALLGOALS (ASM_SIMP_TAC add_ss));
+{\out Level 2}
+{\out m + Suc(n) = Suc(m + n)}
+{\out No subgoals!}
+\end{ttbox}
+Some goals contain free function variables.  The simplifier must have
+congruence rules for those function variables, or it will be unable to
+simplify their arguments:
+\begin{ttbox}
+val f_congs = mk_typed_congs Nat.thy [("f","nat => nat")];
+val f_ss = add_ss addcongs f_congs;
+prths f_congs;
+{\out ?Xa = ?Ya ==> f(?Xa) = f(?Ya)}
+\end{ttbox}
+Here is a conjecture to be proved for an arbitrary function~$f$ satisfying
+the law $f(Suc(n)) = Suc(f(n))$:
+\begin{ttbox}
+val [prem] = goal Nat.thy
+    "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
+{\out Level 0}
+{\out f(i + j) = i + f(j)}
+{\out  1. f(i + j) = i + f(j)}
+\ttbreak
+by (res_inst_tac [("n","i")] induct 1);
+{\out Level 1}
+{\out f(i + j) = i + f(j)}
+{\out  1. f(0 + j) = 0 + f(j)}
+{\out  2. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)}
+\end{ttbox}
+We simplify each subgoal in turn.  The first one is trivial:
+\begin{ttbox}
+by (SIMP_TAC f_ss 1);
+{\out Level 2}
+{\out f(i + j) = i + f(j)}
+{\out  1. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)}
+\end{ttbox}
+The remaining subgoal requires rewriting by the premise, shown
+below, so we add it to {\tt f_ss}:
+\begin{ttbox}
+prth prem;
+{\out f(Suc(?n)) = Suc(f(?n))  [!!n. f(Suc(n)) = Suc(f(n))]}
+by (ASM_SIMP_TAC (f_ss addrews [prem]) 1);
+{\out Level 3}
+{\out f(i + j) = i + f(j)}
+{\out No subgoals!}
+\end{ttbox}
+
+
+\section{Setting up the simplifier} \label{SimpFun-input}
+\index{simplification!setting up|bold}
+To set up a simplifier for a new logic, the \ML\ functor
+\ttindex{SimpFun} needs to be supplied with theorems to justify
+rewriting.  A rewrite relation must be reflexive and transitive; symmetry
+is not necessary.  Hence the package is also applicable to non-symmetric
+relations such as occur in operational semantics.  In the sequel, $\gg$
+denotes some {\bf reduction relation}: a binary relation to be used for
+rewriting.  Several reduction relations can be used at once.  In {\tt FOL},
+both $=$ (on terms) and $\bimp$ (on formulae) can be used for rewriting.
+
+The argument to {\tt SimpFun} is a structure with signature
+\ttindexbold{SIMP_DATA}:
+\begin{ttbox}
+signature SIMP_DATA =
+sig
+  val case_splits  : (thm * string) list
+  val dest_red     : term -> term * term * term
+  val mk_rew_rules : thm -> thm list
+  val norm_thms    : (thm*thm) list
+  val red1         : thm
+  val red2         : thm 
+  val refl_thms    : thm list
+  val subst_thms   : thm list 
+  val trans_thms   : thm list
+end;
+\end{ttbox}
+The components of {\tt SIMP_DATA} need to be instantiated as follows.  Many
+of these components are lists, and can be empty.
+\begin{description}
+\item[\ttindexbold{refl_thms}] 
+supplies reflexivity theorems of the form $\Var{x} \gg
+\Var{x}$.  They must not have additional premises as, for example,
+$\Var{x}\in\Var{A} \Imp \Var{x} = \Var{x}\in\Var{A}$ in type theory.
+
+\item[\ttindexbold{trans_thms}] 
+supplies transitivity theorems of the form
+$\List{\Var{x}\gg\Var{y}; \Var{y}\gg\Var{z}} \Imp {\Var{x}\gg\Var{z}}$.
+
+\item[\ttindexbold{red1}] 
+is a theorem of the form $\List{\Var{P}\gg\Var{Q};
+\Var{P}} \Imp \Var{Q}$, where $\gg$ is a relation between formulae, such as
+$\bimp$ in {\tt FOL}.
+
+\item[\ttindexbold{red2}] 
+is a theorem of the form $\List{\Var{P}\gg\Var{Q};
+\Var{Q}} \Imp \Var{P}$, where $\gg$ is a relation between formulae, such as
+$\bimp$ in {\tt FOL}.
+
+\item[\ttindexbold{mk_rew_rules}] 
+is a function that extracts rewrite rules from theorems.  A rewrite rule is
+a theorem of the form $\List{\ldots}\Imp s \gg t$.  In its simplest form,
+{\tt mk_rew_rules} maps a theorem $t$ to the singleton list $[t]$.  More
+sophisticated versions may do things like
+\[
+\begin{array}{l@{}r@{\quad\mapsto\quad}l}
+\mbox{create formula rewrites:}& P & [P \bimp True] \\[.5ex]
+\mbox{remove negations:}& \lnot P & [P \bimp False] \\[.5ex]
+\mbox{create conditional rewrites:}& P\imp s\gg t & [P\Imp s\gg t] \\[.5ex]
+\mbox{break up conjunctions:}& 
+        (s@1 \gg@1 t@1) \conj (s@2 \gg@2 t@2) & [s@1 \gg@1 t@1, s@2 \gg@2 t@2]
+\end{array}
+\]
+The more theorems are turned into rewrite rules, the better.  The function
+is used in two places:
+\begin{itemize}
+\item 
+$ss$~\ttindex{addrews}~$thms$ applies {\tt mk_rew_rules} to each element of
+$thms$ before adding it to $ss$.
+\item 
+simplification with assumptions (as in \ttindex{ASM_SIMP_TAC}) uses
+{\tt mk_rew_rules} to turn assumptions into rewrite rules.
+\end{itemize}
+
+\item[\ttindexbold{case_splits}] 
+supplies expansion rules for case splits.  The simplifier is designed
+for rules roughly of the kind
+\[ \Var{P}(if(\Var{Q},\Var{x},\Var{y})) \bimp (\Var{Q} \imp \Var{P}(\Var{x}))
+\conj (\lnot\Var{Q} \imp \Var{P}(\Var{y})) 
+\] 
+but is insensitive to the form of the right-hand side.  Other examples
+include product types, where $split ::
+(\alpha\To\beta\To\gamma)\To\alpha*\beta\To\gamma$:
+\[ \Var{P}(split(\Var{f},\Var{p})) \bimp (\forall a~b. \Var{p} =
+{<}a,b{>} \imp \Var{P}(\Var{f}(a,b))) 
+\] 
+Each theorem in the list is paired with the name of the constant being
+eliminated, {\tt"if"} and {\tt"split"} in the examples above.
+
+\item[\ttindexbold{norm_thms}] 
+supports an optimization.  It should be a list of pairs of rules of the
+form $\Var{x} \gg norm(\Var{x})$ and $norm(\Var{x}) \gg \Var{x}$.  These
+introduce and eliminate {\tt norm}, an arbitrary function that should be
+used nowhere else.  This function serves to tag subterms that are in normal
+form.  Such rules can speed up rewriting significantly!
+
+\item[\ttindexbold{subst_thms}] 
+supplies substitution rules of the form
+\[ \List{\Var{x} \gg \Var{y}; \Var{P}(\Var{x})} \Imp \Var{P}(\Var{y}) \]
+They are used to derive congruence rules via \ttindex{mk_congs} and
+\ttindex{mk_typed_congs}.  If $f :: [\tau@1,\cdots,\tau@n]\To\tau$ is a
+constant or free variable, the computation of a congruence rule
+\[\List{\Var{x@1} \gg@1 \Var{y@1}; \ldots; \Var{x@n} \gg@n \Var{y@n}}
+\Imp f(\Var{x@1},\ldots,\Var{x@n}) \gg f(\Var{y@1},\ldots,\Var{y@n}) \]
+requires a reflexivity theorem for some reduction ${\gg} ::
+\alpha\To\alpha\To\sigma$ such that $\tau$ is an instance of $\alpha$.  If a
+suitable reflexivity law is missing, no congruence rule for $f$ can be
+generated.   Otherwise an $n$-ary congruence rule of the form shown above is
+derived, subject to the availability of suitable substitution laws for each
+argument position.  
+
+A substitution law is suitable for argument $i$ if it
+uses a reduction ${\gg@i} :: \alpha@i\To\alpha@i\To\sigma@i$ such that
+$\tau@i$ is an instance of $\alpha@i$.  If a suitable substitution law for
+argument $i$ is missing, the $i^{th}$ premise of the above congruence rule
+cannot be generated and hence argument $i$ cannot be rewritten.  In the
+worst case, if there are no suitable substitution laws at all, the derived
+congruence simply degenerates into a reflexivity law.
+
+\item[\ttindexbold{dest_red}] 
+takes reductions apart.  Given a term $t$ representing the judgement
+\mbox{$a \gg b$}, \verb$dest_red$~$t$ should return a triple $(c,ta,tb)$
+where $ta$ and $tb$ represent $a$ and $b$, and $c$ is a term of the form
+\verb$Const(_,_)$, the reduction constant $\gg$.  
+
+Suppose the logic has a coercion function like $Trueprop::o\To prop$, as do
+{\tt FOL} and~{\tt HOL}\@.  If $\gg$ is a binary operator (not necessarily
+infix), the following definition does the job:
+\begin{verbatim}
+   fun dest_red( _ $ (c $ ta $ tb) ) = (c,ta,tb);
+\end{verbatim}
+The wildcard pattern {\tt_} matches the coercion function.
+\end{description}
+
+
+\section{A sample instantiation}
+Here is the instantiation of {\tt SIMP_DATA} for {\FOL}.  The code for {\tt
+mk_rew_rules} is not shown; see the file {\tt FOL/simpdata.ML}.
+\begin{ttbox}
+structure FOL_SimpData : SIMP_DATA =
+  struct
+  val refl_thms      = [ \(\Var{x}=\Var{x}\), \(\Var{P}\bimp\Var{P}\) ]
+  val trans_thms     = [ \(\List{\Var{x}=\Var{y};\Var{y}=\Var{z}}\Imp\Var{x}=\Var{z}\),
+                         \(\List{\Var{P}\bimp\Var{Q};\Var{Q}\bimp\Var{R}}\Imp\Var{P}\bimp\Var{R}\) ]
+  val red1           = \(\List{\Var{P}\bimp\Var{Q}; \Var{P}} \Imp \Var{Q}\)
+  val red2           = \(\List{\Var{P}\bimp\Var{Q}; \Var{Q}} \Imp \Var{P}\)
+  val mk_rew_rules   = ...
+  val case_splits    = [ \(\Var{P}(if(\Var{Q},\Var{x},\Var{y})) \bimp\)
+                           \((\Var{Q} \imp \Var{P}(\Var{x})) \conj (\lnot\Var{Q} \imp \Var{P}(\Var{y}))\) ]
+  val norm_thms      = [ (\(\Var{x}=norm(\Var{x})\),\(norm(\Var{x})=\Var{x}\)),
+                        (\(\Var{P}\bimp{}NORM(\Var{P}\)), \(NORM(\Var{P})\bimp\Var{P}\)) ]
+  val subst_thms     = [ \(\List{\Var{x}=\Var{y}; \Var{P}(\Var{x})}\Imp\Var{P}(\Var{y})\) ]
+  val dest_red       = fn (_ $ (opn $ lhs $ rhs)) => (opn,lhs,rhs)
+  end;
+\end{ttbox}
+
+\index{simplification|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/simplifier-eg.txt	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,73 @@
+> goal Nat.thy "m+0 = m";
+Level 0
+m + 0 = m
+ 1. m + 0 = m
+> by (res_inst_tac [("n","m")] induct 1);
+Level 1
+m + 0 = m
+ 1. 0 + 0 = 0
+ 2. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)
+> by (simp_tac add_ss 1);
+Level 2
+m + 0 = m
+ 1. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)
+> by (asm_simp_tac add_ss 1);
+Level 3
+m + 0 = m
+No subgoals!
+
+
+> goal Nat.thy "m+Suc(n) = Suc(m+n)";
+Level 0
+m + Suc(n) = Suc(m + n)
+ 1. m + Suc(n) = Suc(m + n)
+val it = [] : thm list   
+> by (res_inst_tac [("n","m")] induct 1);
+Level 1
+m + Suc(n) = Suc(m + n)
+ 1. 0 + Suc(n) = Suc(0 + n)
+ 2. !!x. x + Suc(n) = Suc(x + n) ==> Suc(x) + Suc(n) = Suc(Suc(x) + n)
+val it = () : unit   
+> by (simp_tac add_ss 1);
+Level 2
+m + Suc(n) = Suc(m + n)
+ 1. !!x. x + Suc(n) = Suc(x + n) ==> Suc(x) + Suc(n) = Suc(Suc(x) + n)
+val it = () : unit   
+> trace_simp := true;
+val it = () : unit   
+> by (asm_simp_tac add_ss 1);
+Rewriting:
+Suc(x) + Suc(n) == Suc(x + Suc(n))
+Rewriting:
+x + Suc(n) == Suc(x + n)
+Rewriting:
+Suc(x) + n == Suc(x + n)
+Rewriting:
+Suc(Suc(x + n)) = Suc(Suc(x + n)) == True
+Level 3
+m + Suc(n) = Suc(m + n)
+No subgoals!
+val it = () : unit   
+
+
+
+> val prems = goal Nat.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
+Level 0
+f(i + j) = i + f(j)
+ 1. f(i + j) = i + f(j)
+> prths prems;
+f(Suc(?n)) = Suc(f(?n))  [!!n. f(Suc(n)) = Suc(f(n))]
+
+> by (res_inst_tac [("n","i")] induct 1);
+Level 1
+f(i + j) = i + f(j)
+ 1. f(0 + j) = 0 + f(j)
+ 2. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)
+> by (simp_tac f_ss 1);
+Level 2
+f(i + j) = i + f(j)
+ 1. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)
+> by (asm_simp_tac (f_ss addrews prems) 1);
+Level 3
+f(i + j) = i + f(j)
+No subgoals!
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/simplifier.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,360 @@
+%% $Id$
+\chapter{Simplification} \label{simp-chap}
+\index{simplification|(}
+
+
+This chapter describes Isabelle's generic simplification package, which
+provides a suite of simplification tactics.  This rewriting package is less
+general than its predecessor --- it works only for the equality relation,
+not arbitrary preorders --- but it is fast and flexible.  It performs
+conditional and unconditional rewriting and uses contextual information
+(``local assumptions'').  It provides a few general hooks, which can
+provide automatic case splits during rewriting, for example.  The
+simplifier is set up for many of Isabelle's logics: {\tt FOL}, {\tt ZF},
+{\tt LCF} and {\tt HOL}.
+
+
+\section{Simplification sets}
+\index{simplification sets} 
+
+The simplification tactics are controlled by {\bf simpsets}.  These consist
+of five components: rewrite rules, congruence rules, the subgoaler, the
+solver and the looper.  Normally, the simplifier is set up with sensible
+defaults, so that most simplifier calls specify only rewrite rules.
+Sophisticated usage of the other components can be highly effective, but
+most users should never worry about them.
+
+\subsection{Rewrite rules}\index{rewrite rules}
+
+Rewrite rules are theorems like $Suc(\Var{m}) + \Var{n} = \Var{m} +
+Suc(\Var{n})$, $\Var{P}\conj\Var{P} \bimp \Var{P}$, or $\Var{A} \un \Var{B}
+\equiv \{x.x\in A \disj x\in B\}$.  {\bf Conditional} rewrites such as
+$\Var{m}<\Var{n} \Imp \Var{m}/\Var{n} = 0$ are permitted; the conditions
+can be arbitrary terms.  The infix operation \ttindex{addsimps} adds new
+rewrite rules, while \ttindex{delsimps} deletes rewrite rules from a
+simpset.
+
+Theorems added via \ttindex{addsimps} need not be equalities to start with.
+Each simpset contains a (user-definable) function for extracting equalities
+from arbitrary theorems.  For example $\neg(x\in \{\})$ could be turned
+into $x\in \{\} \equiv False$.  This function can be set with
+\ttindex{setmksimps} but only the definer of a logic should need to do
+this.  Exceptionally, one may want to install a selective version of
+\ttindex{mksimps} in order to filter out looping rewrite rules arising from
+local assumptions (see below).
+
+Internally, all rewrite rules are translated into meta-equalities:
+theorems with conclusion $lhs \equiv rhs$.  To this end every simpset contains
+a function of type \verb$thm -> thm list$ to extract a list
+of meta-equalities from a given theorem.
+
+\begin{warn}\index{rewrite rules}
+  The left-hand side of a rewrite rule must look like a first-order term:
+  after eta-contraction, none of its unknowns should have arguments.  Hence
+  ${\Var{i}+(\Var{j}+\Var{k})} = {(\Var{i}+\Var{j})+\Var{k}}$ and $\neg(\forall
+  x.\Var{P}(x)) \bimp (\exists x.\neg\Var{P}(x))$ are acceptable, whereas
+  $\Var{f}(\Var{x})\in {\tt range}(\Var{f}) = True$ is not.  However, you can
+  replace the offending subterms by new variables and conditions: $\Var{y} =
+  \Var{f}(\Var{x}) \Imp \Var{y}\in {\tt range}(\Var{f}) = True$ is again
+  acceptable.
+\end{warn}
+
+\subsection {Congruence rules}\index{congruence rules}
+Congruence rules are meta-equalities of the form
+\[ \List{\dots} \Imp
+   f(\Var{x@1},\ldots,\Var{x@n}) \equiv f(\Var{y@1},\ldots,\Var{y@n}).
+\]
+They control the simplification of the arguments of certain constants.  For
+example, some arguments can be simplified under additional assumptions:
+\[ \List{\Var{P@1} \bimp \Var{Q@1};\; \Var{Q@1} \Imp \Var{P@2} \bimp \Var{Q@2}}
+   \Imp (\Var{P@1} \imp \Var{P@2}) \equiv (\Var{Q@1} \imp \Var{Q@2})
+\]
+This rule assumes $Q@1$ and any rewrite rules it implies, while
+simplifying~$P@2$.  Such ``local'' assumptions are effective for rewriting
+formulae such as $x=0\imp y+x=y$.  The next example makes similar use of
+such contextual information in bounded quantifiers:
+\[ \List{\Var{A}=\Var{B};\; \Forall x. x\in \Var{B} \Imp \Var{P}(x) = \Var{Q}(x)}
+   \Imp (\forall x\in \Var{A}.\Var{P}(x)) = (\forall x\in \Var{B}.\Var{Q}(x))
+\]
+This congruence rule supplies contextual information for simplifying the
+arms of a conditional expressions:
+\[ \List{\Var{p}=\Var{q};~ \Var{q} \Imp \Var{a}=\Var{c};~
+         \neg\Var{q} \Imp \Var{b}=\Var{d}} \Imp
+   if(\Var{p},\Var{a},\Var{b}) \equiv if(\Var{q},\Var{c},\Var{d})
+\]
+
+A congruence rule can also suppress simplification of certain arguments.
+Here is an alternative congruence rule for conditional expressions:
+\[ \Var{p}=\Var{q} \Imp
+   if(\Var{p},\Var{a},\Var{b}) \equiv if(\Var{q},\Var{a},\Var{b})
+\]
+Only the first argument is simplified; the others remain unchanged.
+This can make simplification much faster, but may require an extra case split
+to prove the goal.  
+
+Congruence rules are added using \ttindex{addeqcongs}.  Their conclusion
+must be a meta-equality, as in the examples above.  It is more
+natural to derive the rules with object-logic equality, for example
+\[ \List{\Var{P@1} \bimp \Var{Q@1};\; \Var{Q@1} \Imp \Var{P@2} \bimp \Var{Q@2}}
+   \Imp (\Var{P@1} \imp \Var{P@2}) \bimp (\Var{Q@1} \imp \Var{Q@2}),
+\]
+So each object-logic should provide an alternative combinator
+\ttindex{addcongs} that expects object-equalities and translates them into
+meta-equalities.
+
+\subsection{The subgoaler} \index{subgoaler}
+The subgoaler is the tactic used to solve subgoals arising out of
+conditional rewrite rules or congruence rules.  The default should be
+simplification itself.  Occasionally this strategy needs to be changed.  For
+example, if the premise of a conditional rule is an instance of its
+conclusion, as in $Suc(\Var{m}) < \Var{n} \Imp \Var{m} < \Var{n}$, the
+default strategy could loop.
+
+The subgoaler can be set explicitly with \ttindex{setsubgoaler}.  For
+example, the subgoaler
+\begin{ttbox}
+fun subgoal_tac ss = resolve_tac (prems_of_ss ss) ORELSE' 
+                     asm_simp_tac ss;
+\end{ttbox}
+tries to solve the subgoal with one of the premises and calls
+simplification only if that fails; here {\tt prems_of_ss} extracts the
+current premises from a simpset.
+
+\subsection{The solver} \index{solver}
+The solver attempts to solve a subgoal after simplification, say by
+recognizing it as a tautology. It can be set with \ttindex{setsolver}.
+Occasionally, simplification on its own is not enough to reduce the subgoal
+to a tautology; additional means, like \verb$fast_tac$, may be necessary.
+
+\begin{warn}
+  Rewriting does not instantiate unknowns.  Trying to rewrite $a\in
+  \Var{A}$ with the rule $\Var{x}\in \{\Var{x}\}$ leads nowhere.  The
+  solver, however, is an arbitrary tactic and may instantiate unknowns as
+  it pleases.  This is the only way the simplifier can handle a conditional
+  rewrite rule whose condition contains extra variables.
+\end{warn}
+
+\begin{warn}
+  If you want to supply your own subgoaler or solver, read on.  The subgoaler
+  is also used to solve the premises of congruence rules, which are usually
+  of the form $s = \Var{x}$, where $s$ needs to be simplified and $\Var{x}$
+  needs to be instantiated with the result. Hence the subgoaler should call
+  the simplifier at some point. The simplifier will then call the solver,
+  which must therefore be prepared to solve goals of the form $t = \Var{x}$,
+  usually by reflexivity. In particular, reflexivity should be tried before
+  any of the fancy tactics like {\tt fast_tac}. It may even happen that, due
+  to simplification, the subgoal is no longer an equality. For example $False
+  \bimp \Var{Q}$ could be rewritten to $\neg\Var{Q}$, in which case the
+  solver must also try resolving with the theorem $\neg False$.
+
+  If the simplifier aborts with the message {\tt Failed congruence proof!},
+  it is due to the subgoaler or solver who failed to prove a premise of a
+  congruence rule.
+\end{warn}
+
+\subsection{The looper} \index{looper}
+The looper is a tactic that is applied after simplification, in case the
+solver failed to solve the simplified goal.  If the looper succeeds, the
+simplification process is started all over again.  Each of the subgoals
+generated by the looper is attacked in turn, in reverse order.  A
+typical looper is case splitting: the expansion of a conditional.  Another
+possibility is to apply an elimination rule on the assumptions.  More
+adventurous loopers could start an induction.  The looper is set with 
+\ttindex{setloop}.
+
+
+\begin{figure}
+\indexbold{*SIMPLIFIER}
+\begin{ttbox}
+infix addsimps addeqcongs delsimps
+      setsubgoaler setsolver setloop setmksimps;
+
+signature SIMPLIFIER =
+sig
+  type simpset
+  val simp_tac:          simpset -> int -> tactic
+  val asm_simp_tac:      simpset -> int -> tactic
+  val asm_full_simp_tac: simpset -> int -> tactic
+  val addeqcongs:  simpset * thm list -> simpset
+  val addsimps:    simpset * thm list -> simpset
+  val delsimps:    simpset * thm list -> simpset
+  val empty_ss:     simpset
+  val merge_ss:     simpset * simpset -> simpset
+  val setsubgoaler: simpset * (simpset -> int -> tactic) -> simpset
+  val setsolver:    simpset * (thm list -> int -> tactic) -> simpset
+  val setloop:      simpset * (int -> tactic) -> simpset
+  val setmksimps:   simpset * (thm -> thm list) -> simpset
+  val prems_of_ss:  simpset -> thm list
+  val rep_ss:       simpset -> \{simps: thm list, congs: thm list\}
+end;
+\end{ttbox}
+\caption{The signature \ttindex{SIMPLIFIER}} \label{SIMPLIFIER}
+\end{figure}
+
+
+\section{The simplification tactics} \label{simp-tactics}
+\index{simplification!tactics|bold}
+\index{tactics!simplification|bold}
+
+The actual simplification work is performed by the following tactics.  The
+rewriting strategy is strictly bottom up, except for congruence rules, which
+are applied while descending into a term.  Conditions in conditional rewrite
+rules are solved recursively before the rewrite rule is applied.
+
+There are three basic simplification tactics:
+\begin{description}
+\item[\ttindexbold{simp_tac} $ss$ $i$] simplifies subgoal~$i$ using the rules
+  in~$ss$.  It may solve the subgoal completely if it has become trivial,
+  using the solver.
+  
+\item[\ttindexbold{asm_simp_tac}] is like \verb$simp_tac$, but also uses
+  assumptions as additional rewrite rules.
+
+\item[\ttindexbold{asm_full_simp_tac}] is like \verb$asm_simp_tac$, but also
+  simplifies the assumptions one by one, using each assumption in the
+  simplification of the following ones.
+\end{description}
+Using the simplifier effectively may take a bit of experimentation.  The
+tactics can be traced with the ML command \verb$trace_simp := true$.  To
+remind yourself of what is in a simpset, use the function \verb$rep_ss$ to
+return its simplification and congruence rules.
+
+\section{Example: using the simplifier}
+\index{simplification!example}
+Assume we are working within {\tt FOL} and that
+\begin{description}
+\item[\tt Nat.thy] is a theory including the constants $0$, $Suc$ and $+$,
+\item[\tt add_0] is the rewrite rule $0+n = n$,
+\item[\tt add_Suc] is the rewrite rule $Suc(m)+n = Suc(m+n)$,
+\item[\tt induct] is the induction rule
+$\List{P(0); \Forall x. P(x)\Imp P(Suc(x))} \Imp P(n)$.
+\item[\tt FOL_ss] is a basic simpset for {\tt FOL}.\footnote
+{These examples reside on the file {\tt FOL/ex/nat.ML}.} 
+\end{description}
+
+We create a simpset for natural numbers by extending~{\tt FOL_ss}:
+\begin{ttbox}
+val add_ss = FOL_ss addsimps [add_0, add_Suc];
+\end{ttbox}
+Proofs by induction typically involve simplification:
+\begin{ttbox}
+goal Nat.thy "m+0 = m";
+{\out Level 0}
+{\out m + 0 = m}
+{\out  1. m + 0 = m}
+\ttbreak
+by (res_inst_tac [("n","m")] induct 1);
+{\out Level 1}
+{\out m + 0 = m}
+{\out  1. 0 + 0 = 0}
+{\out  2. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)}
+\end{ttbox}
+Simplification solves the first subgoal:
+\begin{ttbox}
+by (simp_tac add_ss 1);
+{\out Level 2}
+{\out m + 0 = m}
+{\out  1. !!x. x + 0 = x ==> Suc(x) + 0 = Suc(x)}
+\end{ttbox}
+The remaining subgoal requires \ttindex{asm_simp_tac} in order to use the
+induction hypothesis as a rewrite rule:
+\begin{ttbox}
+by (asm_simp_tac add_ss 1);
+{\out Level 3}
+{\out m + 0 = m}
+{\out No subgoals!}
+\end{ttbox}
+
+\medskip
+The next proof is similar.
+\begin{ttbox}
+goal Nat.thy "m+Suc(n) = Suc(m+n)";
+{\out Level 0}
+{\out m + Suc(n) = Suc(m + n)}
+{\out  1. m + Suc(n) = Suc(m + n)}
+\ttbreak
+by (res_inst_tac [("n","m")] induct 1);
+{\out Level 1}
+{\out m + Suc(n) = Suc(m + n)}
+{\out  1. 0 + Suc(n) = Suc(0 + n)}
+{\out  2. !!x. x + Suc(n) = Suc(x + n) ==> Suc(x) + Suc(n) = Suc(Suc(x) + n)}
+\ttbreak
+by (simp_tac add_ss 1);
+{\out Level 2}
+{\out m + Suc(n) = Suc(m + n)}
+{\out  1. !!x. x + Suc(n) = Suc(x + n) ==> Suc(x) + Suc(n) = Suc(Suc(x) + n)}
+\end{ttbox}
+Switching tracing on illustrates how the simplifier solves the remaining
+subgoal: 
+\begin{ttbox}
+trace_simp := true;
+by (asm_simp_tac add_ss 1);
+{\out Rewriting:}
+{\out Suc(x) + Suc(n) == Suc(x + Suc(n))}
+{\out Rewriting:}
+{\out x + Suc(n) == Suc(x + n)}
+{\out Rewriting:}
+{\out Suc(x) + n == Suc(x + n)}
+{\out Rewriting:}
+{\out Suc(Suc(x + n)) = Suc(Suc(x + n)) == True}
+{\out Level 3}
+{\out m + Suc(n) = Suc(m + n)}
+{\out No subgoals!}
+\end{ttbox}
+As usual, many variations are possible.  At Level~1 we could have solved
+both subgoals at once using the tactical \ttindex{ALLGOALS}:
+\begin{ttbox}
+by (ALLGOALS (asm_simp_tac add_ss));
+{\out Level 2}
+{\out m + Suc(n) = Suc(m + n)}
+{\out No subgoals!}
+\end{ttbox}
+
+\medskip
+Here is a conjecture to be proved for an arbitrary function~$f$ satisfying
+the law $f(Suc(n)) = Suc(f(n))$:\footnote{The previous
+  simplifier required congruence rules for such function variables in
+  order to simplify their arguments.  The present simplifier can be given
+  congruence rules to realize non-standard simplification of a function's
+  arguments, but this is seldom necessary.}
+\begin{ttbox}
+val [prem] = goal Nat.thy
+    "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
+{\out Level 0}
+{\out f(i + j) = i + f(j)}
+{\out  1. f(i + j) = i + f(j)}
+{\out val prem = "f(Suc(?n)) = Suc(f(?n))  [!!n. f(Suc(n)) = Suc(f(n))]" : thm}
+\ttbreak
+by (res_inst_tac [("n","i")] induct 1);
+{\out Level 1}
+{\out f(i + j) = i + f(j)}
+{\out  1. f(0 + j) = 0 + f(j)}
+{\out  2. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)}
+\end{ttbox}
+We simplify each subgoal in turn.  The first one is trivial:
+\begin{ttbox}
+by (simp_tac add_ss 1);
+{\out Level 2}
+{\out f(i + j) = i + f(j)}
+{\out  1. !!x. f(x + j) = x + f(j) ==> f(Suc(x) + j) = Suc(x) + f(j)}
+\end{ttbox}
+The remaining subgoal requires rewriting by the premise, so we add it to
+{\tt add_ss}: 
+\begin{ttbox}
+by (asm_simp_tac (add_ss addsimps [prem]) 1);
+{\out Level 3}
+{\out f(i + j) = i + f(j)}
+{\out No subgoals!}
+\end{ttbox}
+
+No documentation is available on setting up the simplifier for new logics.
+Please consult {\tt FOL/simpdata.ML} to see how this is done, and {\tt
+  FOL/simpdata.ML} for a fairly sophisticated translation of formulae into
+rewrite rules.
+
+%%\section{Setting up the simplifier} \label{SimpFun-input}
+%%Should be written!
+
+
+\index{simplification|)}
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/substitution.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,165 @@
+%% $Id$
+\chapter{Substitution Tactics} \label{substitution}
+\index{substitution|(}
+Replacing equals by equals is a basic form of reasoning.  Isabelle supports
+several kinds of equality reasoning.  {\bf Substitution} means to replace
+free occurrences of~$t$ by~$u$ in a subgoal.  This is easily done, given an
+equality $t=u$, provided the logic possesses the appropriate rule ---
+unless you want to substitute even in the assumptions.  The tactic
+\ttindex{hyp_subst_tac} performs substitution in the assumptions, but it
+works via object-level implication, and therefore must be specially set up
+for each suitable object-logic.
+
+Substitution should not be confused with object-level {\bf rewriting}.
+Given equalities of the form $t=u$, rewriting replaces instances of~$t$ by
+corresponding instances of~$u$, and continues until it reaches a normal
+form.  Substitution handles `one-off' replacements by particular
+equalities, while rewriting handles general equalities.
+Chapter~\ref{simp-chap} discusses Isabelle's rewriting tactics.
+
+
+\section{Simple substitution}
+\index{substitution!simple}
+Many logics include a substitution rule of the form\indexbold{*subst}
+$$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{a})} \Imp 
+   \Var{P}(\Var{b})  \eqno(subst)$$
+In backward proof, this may seem difficult to use: the conclusion
+$\Var{P}(\Var{b})$ admits far too many unifiers.  But, if the theorem {\tt
+eqth} asserts $t=u$, then \hbox{\tt eqth RS subst} is the derived rule
+\[ \Var{P}(t) \Imp \Var{P}(u). \]
+Provided $u$ is not an unknown, resolution with this rule is
+well-behaved.\footnote{Unifying $\Var{P}(u)$ with a formula~$Q$
+expresses~$Q$ in terms of its dependence upon~$u$.  There are still $2^k$
+unifiers, if $Q$ has $k$ occurrences of~$u$, but Isabelle ensures that
+the first unifier includes all the occurrences.}  To replace $u$ by~$t$ in
+subgoal~$i$, use
+\begin{ttbox} 
+resolve_tac [eqth RS subst] \(i\) {\it.}
+\end{ttbox}
+To replace $t$ by~$u$ in
+subgoal~$i$, use
+\begin{ttbox} 
+resolve_tac [eqth RS ssubst] \(i\) {\it,}
+\end{ttbox}
+where \ttindexbold{ssubst} is the `swapped' substitution rule
+$$ \List{\Var{a}=\Var{b}; \Var{P}(\Var{b})} \Imp 
+   \Var{P}(\Var{a}).  \eqno(ssubst)$$
+If \ttindex{sym} denotes the symmetry rule
+\(\Var{a}=\Var{b}\Imp\Var{b}=\Var{a}\), then {\tt ssubst} is just
+\hbox{\tt sym RS subst}.  Many logics with equality include the rules {\tt
+subst} and {\tt ssubst}, as well as {\tt refl}, {\tt sym} and {\tt trans}
+(for the usual equality laws).
+
+Elim-resolution is well-behaved with assumptions of the form $t=u$.
+To replace $u$ by~$t$ or $t$ by~$u$ in subgoal~$i$, use
+\begin{ttbox} 
+eresolve_tac [subst] \(i\)    {\it or}    eresolve_tac [ssubst] \(i\) {\it.}
+\end{ttbox}
+
+
+\section{Substitution in the hypotheses}
+\index{substitution!in hypotheses}
+Substitution rules, like other rules of natural deduction, do not affect
+the assumptions.  This can be inconvenient.  Consider proving the subgoal
+\[ \List{c=a; c=b} \Imp a=b. \]
+Calling \hbox{\tt eresolve_tac [ssubst] \(i\)} simply discards the
+assumption~$c=a$, since $c$ does not occur in~$a=b$.  Of course, we can
+work out a solution.  First apply \hbox{\tt eresolve_tac [subst] \(i\)},
+replacing~$a$ by~$c$:
+\[ \List{c=b} \Imp c=b \]
+Equality reasoning can be difficult, but this trivial proof requires
+nothing more sophisticated than substitution in the assumptions.
+Object-logics that include the rule~$(subst)$ provide a tactic for this
+purpose:
+\begin{ttbox} 
+hyp_subst_tac : int -> tactic
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{hyp_subst_tac} {\it i}] 
+selects an equality assumption of the form $t=u$ or $u=t$, where $t$ is a
+free variable or parameter.  Deleting this assumption, it replaces $t$
+by~$u$ throughout subgoal~$i$, including the other assumptions.
+\end{description}
+The term being replaced must be a free variable or parameter.  Substitution
+for constants is usually unhelpful, since they may appear in other
+theorems.  For instance, the best way to use the assumption $0=1$ is to
+contradict a theorem that states $0\not=1$, rather than to replace 0 by~1
+in the subgoal!
+
+Replacing a free variable causes similar problems if they appear in the
+premises of a rule being derived --- the substitution affects object-level
+assumptions, not meta-level assumptions.  For instance, replacing~$a$
+by~$b$ could make the premise~$P(a)$ worthless.  To avoid this problem, call
+\ttindex{cut_facts_tac} to insert the atomic premises as object-level
+assumptions.
+
+
+\section{Setting up {\tt hyp_subst_tac}} 
+Many Isabelle object-logics, such as {\tt FOL}, {\tt HOL} and their
+descendants, come with {\tt hyp_subst_tac} already defined.  A few others,
+such as {\tt CTT}, do not support this tactic because they lack the
+rule~$(subst)$.  When defining a new logic that includes a substitution
+rule and implication, you must set up {\tt hyp_subst_tac} yourself.  It
+is packaged as the \ML{} functor \ttindex{HypsubstFun}, which takes the
+argument signature~\ttindexbold{HYPSUBST_DATA}:
+\begin{ttbox} 
+signature HYPSUBST_DATA =
+  sig
+  val subst      : thm
+  val sym        : thm
+  val rev_cut_eq : thm
+  val imp_intr   : thm
+  val rev_mp     : thm
+  val dest_eq    : term -> term*term
+  end;
+\end{ttbox}
+Thus, the functor requires the following items:
+\begin{description}
+\item[\ttindexbold{subst}] should be the substitution rule
+$\List{\Var{a}=\Var{b};\; \Var{P}(\Var{a})} \Imp \Var{P}(\Var{b})$.
+
+\item[\ttindexbold{sym}] should be the symmetry rule
+$\Var{a}=\Var{b}\Imp\Var{b}=\Var{a}$.
+
+\item[\ttindexbold{rev_cut_eq}] should have the form
+$\List{\Var{a}=\Var{b};\; \Var{a}=\Var{b}\Imp\Var{R}} \Imp \Var{R}$.
+
+\item[\ttindexbold{imp_intr}] should be the implies introduction
+rule $(\Var{P}\Imp\Var{Q})\Imp \Var{P}\imp\Var{Q}$.
+
+\item[\ttindexbold{rev_mp}] should be the ``reversed'' implies elimination
+rule $\List{\Var{P};  \;\Var{P}\imp\Var{Q}} \Imp \Var{Q}$.
+
+\item[\ttindexbold{dest_eq}] should return the pair~$(t,u)$ when
+applied to the \ML{} term that represents~$t=u$.  For other terms, it
+should raise exception~\ttindex{Match}.
+\end{description}
+The functor resides in {\tt Provers/hypsubst.ML} on the Isabelle
+distribution directory.  It is not sensitive to the precise formalization
+of the object-logic.  It is not concerned with the names of the equality
+and implication symbols, or the types of formula and terms.  Coding the
+function {\tt dest_eq} requires knowledge of Isabelle's representation of
+terms.  For {\tt FOL} it is defined by
+\begin{ttbox} 
+fun dest_eq (Const("Trueprop",_) $ (Const("op =",_)$t$u)) = (t,u);
+\end{ttbox}
+Here {\tt Trueprop} is the coercion from type'~$o$ to type~$prop$, while
+\hbox{\tt op =} is the internal name of the infix operator~{\tt=}.
+Pattern-matching expresses the function concisely, using wildcards~({\tt_})
+to hide the types.
+
+Given a subgoal of the form
+\[ \List{P@1; \cdots ; t=u; \cdots ; P@n} \Imp Q, \]
+\ttindexbold{hyp_subst_tac} locates a suitable equality
+assumption and moves it to the last position using elim-resolution on {\tt
+rev_cut_eq} (possibly re-orienting it using~{\tt sym}):
+\[ \List{P@1; \cdots ; P@n; t=u} \Imp Q \]
+Using $n$ calls of \hbox{\tt eresolve_tac [rev_mp]}, it creates the subgoal
+\[ \List{t=u} \Imp P@1\imp \cdots \imp P@n \imp Q \]
+By \hbox{\tt eresolve_tac [ssubst]}, it replaces~$t$ by~$u$ throughout:
+\[ P'@1\imp \cdots \imp P'@n \imp Q' \]
+Finally, using $n$ calls of \hbox{\tt resolve_tac [imp_intr]}, it restores
+$P'@1$, \ldots, $P'@n$ as assumptions:
+\[ \List{P'@n; \cdots ; P'@1} \Imp Q' \]
+
+\index{substitution|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/tactic.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,654 @@
+%% $Id$
+\chapter{Tactics} \label{tactics}
+\index{tactics|(}
+Tactics have type \ttindexbold{tactic}.  They are essentially
+functions from theorems to theorem sequences, where the theorems represent
+states of a backward proof.  Tactics seldom need to be coded from scratch,
+as functions; the basic tactics suffice for most proofs.
+
+\section{Resolution and assumption tactics}
+{\bf Resolution} is Isabelle's basic mechanism for refining a subgoal using
+a rule.  {\bf Elim-resolution} is particularly suited for elimination
+rules, while {\bf destruct-resolution} is particularly suited for
+destruction rules.  The {\tt r}, {\tt e}, {\tt d} naming convention is
+maintained for several different kinds of resolution tactics, as well as
+the shortcuts in the subgoal module.
+
+All the tactics in this section act on a subgoal designated by a positive
+integer~$i$.  They fail (by returning the empty sequence) if~$i$ is out of
+range.
+
+\subsection{Resolution tactics}
+\index{tactics!resolution|bold}
+\begin{ttbox} 
+resolve_tac  : thm list -> int -> tactic
+eresolve_tac : thm list -> int -> tactic
+dresolve_tac : thm list -> int -> tactic
+forward_tac  : thm list -> int -> tactic 
+\end{ttbox}
+These perform resolution on a list of theorems, $thms$, representing a list
+of object-rules.  When generating next states, they take each of the rules
+in the order given.  Each rule may yield several next states, or none:
+higher-order resolution may yield multiple resolvents.
+\begin{description}
+\item[\ttindexbold{resolve_tac} {\it thms} {\it i}] 
+refines the proof state using the object-rules, which should normally be
+introduction rules.  It resolves an object-rule's conclusion with
+subgoal~$i$ of the proof state.
+
+\item[\ttindexbold{eresolve_tac} {\it thms} {\it i}] 
+refines the proof state by elim-resolution with the object-rules, which
+should normally be elimination rules.  It resolves with a rule, solves
+its first premise by assumption, and finally {\bf deletes} that assumption
+from any new subgoals.
+
+\item[\ttindexbold{dresolve_tac} {\it thms} {\it i}] 
+performs destruct-resolution with the object-rules, which normally should
+be destruction rules.  This replaces an assumption by the result of
+applying one of the rules.
+
+\item[\ttindexbold{forward_tac}] 
+is like \ttindex{dresolve_tac} except that the selected assumption is not
+deleted.  It applies a rule to an assumption, adding the result as a new
+assumption.
+\end{description}
+
+\subsection{Assumption tactics}
+\index{tactics!assumption|bold}
+\begin{ttbox} 
+assume_tac    : int -> tactic
+eq_assume_tac : int -> tactic
+\end{ttbox} 
+\begin{description}
+\item[\ttindexbold{assume_tac} {\it i}] 
+attempts to solve subgoal~$i$ by assumption.
+
+\item[\ttindexbold{eq_assume_tac}] 
+is like {\tt assume_tac} but does not use unification.  It succeeds (with a
+{\bf unique} next state) if one of the assumptions is identical to the
+subgoal's conclusion.  Since it does not instantiate variables, it cannot
+make other subgoals unprovable.  It is intended to be called from proof
+strategies, not interactively.
+\end{description}
+
+\subsection{Matching tactics} \label{match_tac}
+\index{tactics!matching|bold}
+\begin{ttbox} 
+match_tac  : thm list -> int -> tactic
+ematch_tac : thm list -> int -> tactic
+dmatch_tac : thm list -> int -> tactic
+\end{ttbox}
+These are just like the resolution tactics except that they never
+instantiate unknowns in the proof state.  Flexible subgoals are not updated
+willy-nilly, but are left alone.  Matching --- strictly speaking --- means
+treating the unknowns in the proof state as constants; these tactics merely
+discard unifiers that would update the proof state.
+\begin{description}
+\item[\ttindexbold{match_tac} {\it thms} {\it i}] 
+refines the proof state using the object-rules, matching an object-rule's
+conclusion with subgoal~$i$ of the proof state.
+
+\item[\ttindexbold{ematch_tac}] 
+is like {\tt match_tac}, but performs elim-resolution.
+
+\item[\ttindexbold{dmatch_tac}] 
+is like {\tt match_tac}, but performs destruct-resolution.
+\end{description}
+
+
+\subsection{Resolution with instantiation} \label{res_inst_tac}
+\index{tactics!instantiation|bold}
+\begin{ttbox} 
+res_inst_tac  : (string*string)list -> thm -> int -> tactic
+eres_inst_tac : (string*string)list -> thm -> int -> tactic
+dres_inst_tac : (string*string)list -> thm -> int -> tactic
+forw_inst_tac : (string*string)list -> thm -> int -> tactic
+\end{ttbox}
+These tactics are designed for applying rules such as substitution and
+induction, which cause difficulties for higher-order unification.  The
+tactics accept explicit instantiations for schematic variables in the rule
+--- typically, in the rule's conclusion.  Each instantiation is a pair
+{\tt($v$,$e$)}, where $v$ can be a type variable or ordinary variable:
+\begin{itemize}
+\item If $v$ is the {\bf type variable} {\tt'a}, then
+the rule must contain a schematic type variable \verb$?'a$ of some
+sort~$s$, and $e$ should be a type of sort $s$.
+
+\item If $v$ is the {\bf variable} {\tt P}, then
+the rule must contain a schematic variable \verb$?P$ of some type~$\tau$,
+and $e$ should be a term of some type~$\sigma$ such that $\tau$ and
+$\sigma$ are unifiable.  If the unification of $\tau$ and $\sigma$
+instantiates any schematic type variables in $\tau$, these instantiations
+are recorded for application to the rule.
+\end{itemize}
+Types are instantiated before terms.  Because type instantiations are
+inferred from term instantiations, explicit type instantiations are seldom
+necessary --- if \verb$?t$ has type \verb$?'a$, then the instantiation list
+\verb$[("'a","bool"),("t","True")]$ may be simplified to
+\verb$[("t","True")]$.  Type unknowns in the proof state may cause
+failure because the tactics cannot instantiate them.
+
+The instantiation tactics act on a given subgoal.  Terms in the
+instantiations are type-checked in the context of that subgoal --- in
+particular, they may refer to that subgoal's parameters.  Any unknowns in
+the terms receive subscripts and are lifted over the parameters; thus, you
+may not refer to unknowns in the subgoal.
+
+\begin{description}
+\item[\ttindexbold{res_inst_tac} {\it insts} {\it thm} {\it i}]
+instantiates the rule {\it thm} with the instantiations {\it insts}, as
+described above, and then performs resolution on subgoal~$i$.  Resolution
+typically causes further instantiations; you need not give explicit
+instantiations for every variable in the rule.
+
+\item[\ttindexbold{eres_inst_tac}] 
+is like {\tt res_inst_tac}, but performs elim-resolution.
+
+\item[\ttindexbold{dres_inst_tac}] 
+is like {\tt res_inst_tac}, but performs destruct-resolution.
+
+\item[\ttindexbold{forw_inst_tac}] 
+is like {\tt dres_inst_tac} except that the selected assumption is not
+deleted.  It applies the instantiated rule to an assumption, adding the
+result as a new assumption.
+\end{description}
+
+
+\section{Other basic tactics}
+\subsection{Definitions and meta-level rewriting}
+\index{tactics!meta-rewriting|bold}\index{rewriting!meta-level}
+Definitions in Isabelle have the form $t\equiv u$, where typically $t$ is a
+constant or a constant applied to a list of variables, for example $\it
+sqr(n)\equiv n\times n$.  (Conditional definitions, $\phi\Imp t\equiv u$,
+are not supported.)  {\bf Unfolding} the definition $t\equiv u$ means using
+it as a rewrite rule, replacing~$t$ by~$u$ throughout a theorem.  {\bf
+Folding} $t\equiv u$ means replacing~$u$ by~$t$.  Rewriting continues until
+no rewrites are applicable to any subterm.
+
+There are rules for unfolding and folding definitions; Isabelle does not do
+this automatically.  The corresponding tactics rewrite the proof state,
+yielding a unique result.  See also the {\tt goalw} command, which is the
+easiest way of handling definitions.
+\begin{ttbox} 
+rewrite_goals_tac : thm list -> tactic
+rewrite_tac       : thm list -> tactic
+fold_goals_tac    : thm list -> tactic
+fold_tac          : thm list -> tactic
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{rewrite_goals_tac} {\it defs}]  
+unfolds the {\it defs} throughout the subgoals of the proof state, while
+leaving the main goal unchanged.  Use \ttindex{SELECT_GOAL} to restrict it to a
+particular subgoal.
+
+\item[\ttindexbold{rewrite_tac} {\it defs}]  
+unfolds the {\it defs} throughout the proof state, including the main goal
+--- not normally desirable!
+
+\item[\ttindexbold{fold_goals_tac} {\it defs}]  
+folds the {\it defs} throughout the subgoals of the proof state, while
+leaving the main goal unchanged.
+
+\item[\ttindexbold{fold_tac} {\it defs}]  
+folds the {\it defs} throughout the proof state.
+\end{description}
+
+
+\subsection{Tactic shortcuts}
+\index{shortcuts!for tactics|bold}
+\index{tactics!resolution}\index{tactics!assumption}
+\index{tactics!meta-rewriting}
+\begin{ttbox} 
+rtac     : thm -> int ->tactic
+etac     : thm -> int ->tactic
+dtac     : thm -> int ->tactic
+atac     : int ->tactic
+ares_tac : thm list -> int -> tactic
+rewtac   : thm -> tactic
+\end{ttbox}
+These abbreviate common uses of tactics.
+\begin{description}
+\item[\ttindexbold{rtac} {\it thm} {\it i}] 
+abbreviates \hbox{\tt resolve_tac [{\it thm}] {\it i}}, doing resolution.
+
+\item[\ttindexbold{etac} {\it thm} {\it i}] 
+abbreviates \hbox{\tt eresolve_tac [{\it thm}] {\it i}}, doing elim-resolution.
+
+\item[\ttindexbold{dtac} {\it thm} {\it i}] 
+abbreviates \hbox{\tt dresolve_tac [{\it thm}] {\it i}}, doing
+destruct-resolution.
+
+\item[\ttindexbold{atac} {\it i}] 
+is a synonym for \hbox{\tt assume_tac {\it i}}, doing proof by assumption.
+
+\item[\ttindexbold{ares_tac} {\it thms} {\it i}] 
+tries proof by assumption and resolution; it abbreviates
+\begin{ttbox}
+assume_tac {\it i} ORELSE resolve_tac {\it thms} {\it i}
+\end{ttbox}
+
+\item[\ttindexbold{rewtac} {\it def}] 
+abbreviates \hbox{\tt rewrite_goals_tac [{\it def}]}, unfolding a definition.
+\end{description}
+
+
+\subsection{Inserting premises and facts}\label{cut_facts_tac}
+\index{tactics!for inserting facts|bold}
+\begin{ttbox} 
+cut_facts_tac : thm list -> int -> tactic
+subgoal_tac   :   string -> int -> tactic
+\end{ttbox}
+These tactics add assumptions --- which must be proved, sooner or later ---
+to a given subgoal.
+\begin{description}
+\item[\ttindexbold{cut_facts_tac} {\it thms} {\it i}] 
+  adds the {\it thms} as new assumptions to subgoal~$i$.  Once they have
+  been inserted as assumptions, they become subject to {\tt eresolve_tac}
+  and {\tt rewrite_goals_tac}.  Only rules with no premises are inserted:
+  Isabelle cannot use assumptions that contain $\Imp$ or~$\Forall$.  Sometimes
+  the theorems are premises of a rule being derived, returned by~{\tt
+    goal}; instead of calling this tactic, you could state the goal with an
+  outermost meta-quantifier.
+
+\item[\ttindexbold{subgoal_tac} {\it formula} {\it i}] 
+adds the {\it formula} as a assumption to subgoal~$i$, and inserts the same
+{\it formula} as a new subgoal, $i+1$.
+\end{description}
+
+
+\subsection{Theorems useful with tactics}
+\index{theorems!of pure theory|bold}
+\begin{ttbox} 
+asm_rl: thm 
+cut_rl: thm 
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{asm_rl}] 
+is $\psi\Imp\psi$.  Under elim-resolution it does proof by assumption, and
+\hbox{\tt eresolve_tac (asm_rl::{\it thms}) {\it i}} is equivalent to
+\begin{ttbox} 
+assume_tac {\it i}  ORELSE  eresolve_tac {\it thms} {\it i}
+\end{ttbox}
+
+\item[\ttindexbold{cut_rl}] 
+is $\List{\psi\Imp\theta,\psi}\Imp\theta$.  It is useful for inserting
+assumptions; it underlies \ttindex{forward_tac}, \ttindex{cut_facts_tac}
+and \ttindex{subgoal_tac}.
+\end{description}
+
+
+\section{Obscure tactics}
+\subsection{Tidying the proof state}
+\index{parameters!removing unused|bold}
+\index{flex-flex constraints}
+\begin{ttbox} 
+prune_params_tac : tactic
+flexflex_tac     : tactic
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{prune_params_tac}]  
+  removes unused parameters from all subgoals of the proof state.  It works
+  by rewriting with the theorem $(\Forall x. V)\equiv V$.  This tactic can
+  make the proof state more readable.  It is used with
+  \ttindex{rule_by_tactic} to simplify the resulting theorem.
+
+\item[\ttindexbold{flexflex_tac}]  
+  removes all flex-flex pairs from the proof state by applying the trivial
+  unifier.  This drastic step loses information, and should only be done as
+  the last step of a proof.
+
+  Flex-flex constraints arise from difficult cases of higher-order
+  unification.  To prevent this, use \ttindex{res_inst_tac} to instantiate
+  some variables in a rule~(\S\ref{res_inst_tac}).  Normally flex-flex
+  constraints can be ignored; they often disappear as unknowns get
+  instantiated.
+\end{description}
+
+
+\subsection{Renaming parameters in a goal} \index{parameters!renaming|bold}
+\begin{ttbox} 
+rename_tac        : string -> int -> tactic
+rename_last_tac   : string -> string list -> int -> tactic
+Logic.set_rename_prefix : string -> unit
+Logic.auto_rename       : bool ref      \hfill{\bf initially false}
+\end{ttbox}
+When creating a parameter, Isabelle chooses its name by matching variable
+names via the object-rule.  Given the rule $(\forall I)$ formalized as
+$\left(\Forall x. P(x)\right) \Imp \forall x.P(x)$, Isabelle will note that
+the $\Forall$-bound variable in the premise has the same name as the
+$\forall$-bound variable in the conclusion.  
+
+Sometimes there is insufficient information and Isabelle chooses an
+arbitrary name.  The renaming tactics let you override Isabelle's choice.
+Because renaming parameters has no logical effect on the proof state, the
+{\tt by} command prints the needless message {\tt Warning:\ same as previous
+level}.
+
+Alternatively, you can suppress the naming mechanism described above and
+have Isabelle generate uniform names for parameters.  These names have the
+form $p${\tt a}, $p${\tt b}, $p${\tt c},~\ldots, where $p$ is any desired
+prefix.  They are ugly but predictable.
+
+\begin{description}
+\item[\ttindexbold{rename_tac} {\it str} {\it i}] 
+interprets the string {\it str} as a series of blank-separated variable
+names, and uses them to rename the parameters of subgoal~$i$.  The names
+must be distinct.  If there are fewer names than parameters, then the
+tactic renames the innermost parameters and may modify the remaining ones
+to ensure that all the parameters are distinct.
+
+\item[\ttindexbold{rename_last_tac} {\it prefix} {\it suffixes} {\it i}] 
+generates a list of names by attaching each of the {\it suffixes\/} to the 
+{\it prefix}.  It is intended for coding structural induction tactics,
+where several of the new parameters should have related names.
+
+\item[\ttindexbold{Logic.set_rename_prefix} {\it prefix};] 
+sets the prefix for uniform renaming to~{\it prefix}.  The default prefix
+is {\tt"k"}.
+
+\item[\ttindexbold{Logic.auto_rename} \tt:= true;] 
+makes Isabelle generate uniform names for parameters. 
+\end{description}
+
+
+\subsection{Composition: resolution without lifting}
+\index{tactics!for composition|bold}
+\begin{ttbox}
+compose_tac: (bool * thm * int) -> int -> tactic
+\end{ttbox}
+{\bf Composing} two rules means to resolve them without prior lifting or
+renaming of unknowns.  This low-level operation, which underlies the
+resolution tactics, may occasionally be useful for special effects.
+A typical application is \ttindex{res_inst_tac}, which lifts and instantiates a
+rule, then passes the result to {\tt compose_tac}.
+\begin{description}
+\item[\ttindexbold{compose_tac} ($flag$, $rule$, $m$) $i$] 
+refines subgoal~$i$ using $rule$, without lifting.  The $rule$ is taken to
+have the form $\List{\psi@1; \ldots; \psi@m} \Imp \psi$, where $\psi$ need
+{\bf not} be atomic; thus $m$ determines the number of new subgoals.  If
+$flag$ is {\tt true} then it performs elim-resolution --- it solves the
+first premise of~$rule$ by assumption and deletes that assumption.
+\end{description}
+
+
+\section{Managing lots of rules}
+These operations are not intended for interactive use.  They are concerned
+with the processing of large numbers of rules in automatic proof
+strategies.  Higher-order resolution involving a long list of rules is
+slow.  Filtering techniques can shorten the list of rules given to
+resolution, and can also detect whether a given subgoal is too flexible,
+with too many rules applicable.
+
+\subsection{Combined resolution and elim-resolution} \label{biresolve_tac}
+\index{tactics!resolution}
+\begin{ttbox} 
+biresolve_tac   : (bool*thm)list -> int -> tactic
+bimatch_tac     : (bool*thm)list -> int -> tactic
+subgoals_of_brl : bool*thm -> int
+lessb           : (bool*thm) * (bool*thm) -> bool
+\end{ttbox}
+{\bf Bi-resolution} takes a list of $\it (flag,rule)$ pairs.  For each
+pair, it applies resolution if the flag is~{\tt false} and
+elim-resolution if the flag is~{\tt true}.  A single tactic call handles a
+mixture of introduction and elimination rules.
+
+\begin{description}
+\item[\ttindexbold{biresolve_tac} {\it brls} {\it i}] 
+refines the proof state by resolution or elim-resolution on each rule, as
+indicated by its flag.  It affects subgoal~$i$ of the proof state.
+
+\item[\ttindexbold{bimatch_tac}] 
+is like {\tt biresolve_tac}, but performs matching: unknowns in the
+proof state are never updated (see~\S\ref{match_tac}).
+
+\item[\ttindexbold{subgoals_of_brl}({\it flag},{\it rule})] 
+returns the number of new subgoals that bi-resolution would yield for the
+pair (if applied to a suitable subgoal).  This is $n$ if the flag is
+{\tt false} and $n-1$ if the flag is {\tt true}, where $n$ is the number
+of premises of the rule.  Elim-resolution yields one fewer subgoal than
+ordinary resolution because it solves the major premise by assumption.
+
+\item[\ttindexbold{lessb} ({\it brl1},{\it brl2})] 
+returns the result of 
+\begin{ttbox}
+subgoals_of_brl {\it brl1} < subgoals_of_brl {\it brl2}
+\end{ttbox}
+Note that \hbox{\tt sort lessb {\it brls}} sorts a list of $\it
+(flag,rule)$ pairs by the number of new subgoals they will yield.  Thus,
+those that yield the fewest subgoals should be tried first.
+\end{description}
+
+
+\subsection{Discrimination nets for fast resolution}
+\index{discrimination nets|bold}
+\index{tactics!resolution}
+\begin{ttbox} 
+net_resolve_tac  : thm list -> int -> tactic
+net_match_tac    : thm list -> int -> tactic
+net_biresolve_tac: (bool*thm) list -> int -> tactic
+net_bimatch_tac  : (bool*thm) list -> int -> tactic
+filt_resolve_tac : thm list -> int -> int -> tactic
+could_unify      : term*term->bool
+filter_thms      : (term*term->bool) -> int*term*thm list -> thm list
+\end{ttbox}
+The module \ttindex{Net} implements a discrimination net data structure for
+fast selection of rules \cite[Chapter 14]{charniak80}.  A term is
+classified by the symbol list obtained by flattening it in preorder.
+The flattening takes account of function applications, constants, and free
+and bound variables; it identifies all unknowns and also regards
+$\lambda$-abstractions as unknowns, since they could $\eta$-contract to
+anything.  
+
+A discrimination net serves as a polymorphic dictionary indexed by terms.
+The module provides various functions for inserting and removing items from
+nets.  It provides functions for returning all items whose term could match
+or unify with a target term.  The matching and unification tests are
+overly lax (due to the identifications mentioned above) but they serve as
+useful filters.
+
+A net can store introduction rules indexed by their conclusion, and
+elimination rules indexed by their major premise.  Isabelle provides
+several functions for ``compiling'' long lists of rules into fast
+resolution tactics.  When supplied with a list of theorems, these functions
+build a discrimination net; the net is used when the tactic is applied to a
+goal.  To avoid repreatedly constructing the nets, use currying: bind the
+resulting tactics to \ML{} identifiers.
+
+\begin{description}
+\item[\ttindexbold{net_resolve_tac} {\it thms}] 
+builds a discrimination net to obtain the effect of a similar call to {\tt
+resolve_tac}.
+
+\item[\ttindexbold{net_match_tac} {\it thms}] 
+builds a discrimination net to obtain the effect of a similar call to {\tt
+match_tac}.
+
+\item[\ttindexbold{net_biresolve_tac} {\it brls}] 
+builds a discrimination net to obtain the effect of a similar call to {\tt
+biresolve_tac}.
+
+\item[\ttindexbold{net_bimatch_tac} {\it brls}] 
+builds a discrimination net to obtain the effect of a similar call to {\tt
+bimatch_tac}.
+
+\item[\ttindexbold{filt_resolve_tac} {\it thms} {\it maxr} {\it i}] 
+uses discrimination nets 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}.  
+
+This tactic helps avoid runaway instantiation of unknowns, for example in
+type inference.
+
+\item[\ttindexbold{could_unify} ({\it t},{\it u})] 
+returns {\tt false} if~$t$ and~$u$ are ``obviously'' non-unifiable, and
+otherwise returns~{\tt true}.  It assumes all variables are distinct,
+reporting that {\tt ?a=?a} may unify with {\tt 0=1}.
+
+\item[\ttindexbold{filter_thms} $could\; (limit,prem,thms)$] 
+returns the list of potentially resolvable rules (in {\it thms\/}) for the
+subgoal {\it prem}, using the predicate {\it could\/} to compare the
+conclusion of the subgoal with the conclusion of each rule.  The resulting list
+is no longer than {\it limit}.
+\end{description}
+
+
+\section{Programming tools for proof strategies}
+Do not consider using the primitives discussed in this section unless you
+really need to code tactics from scratch,
+
+\subsection{Operations on type {\tt tactic}}
+\index{tactics!primitives for coding|bold}
+A tactic maps theorems to theorem sequences (lazy lists).  The type
+constructor for sequences is called \ttindex{Sequence.seq}.  To simplify the
+types of tactics and tacticals, Isabelle defines a type of tactics:
+\begin{ttbox} 
+datatype tactic = Tactic of thm -> thm Sequence.seq
+\end{ttbox} 
+{\tt Tactic} and {\tt tapply} convert between tactics and functions.  The
+other operations provide means for coding tactics in a clean style.
+\begin{ttbox} 
+tapply    : tactic * thm -> thm Sequence.seq
+Tactic    :     (thm -> thm Sequence.seq) -> tactic
+PRIMITIVE :                  (thm -> thm) -> tactic  
+STATE     :               (thm -> tactic) -> tactic
+SUBGOAL   : ((term*int) -> tactic) -> int -> tactic
+\end{ttbox} 
+\begin{description}
+\item[\ttindexbold{tapply} {\it tac} {\it thm}]  
+returns the result of applying the tactic, as a function, to {\it thm}.
+
+\item[\ttindexbold{Tactic} {\it f}]  
+packages {\it f} as a tactic.
+
+\item[\ttindexbold{PRIMITIVE} $f$] 
+applies $f$ to the proof state and returns the result as a
+one-element sequence.  This packages the meta-rule~$f$ as a tactic.
+
+\item[\ttindexbold{STATE} $f$] 
+applies $f$ to the proof state and then applies the resulting tactic to the
+same state.  It supports the following style, where the tactic body is
+expressed at a high level, but may peek at the proof state:
+\begin{ttbox} 
+STATE (fn state => {\it some tactic})
+\end{ttbox} 
+
+\item[\ttindexbold{SUBGOAL} $f$ $i$] 
+extracts subgoal~$i$ from the proof state as a term~$t$, and computes a
+tactic by calling~$f(t,i)$.  It applies the resulting tactic to the same
+state.  The tactic body is expressed at a high level, but may peek at a
+particular subgoal:
+\begin{ttbox} 
+SUBGOAL (fn (t,i) => {\it some tactic})
+\end{ttbox} 
+\end{description}
+
+
+\subsection{Tracing}
+\index{tactics!tracing|bold}
+\index{tracing!of tactics}
+\begin{ttbox} 
+pause_tac: tactic
+print_tac: tactic
+\end{ttbox}
+These violate the functional behaviour of tactics by printing information
+when they are applied to a proof state.  Their output may be difficult to
+interpret.  Note that certain of the searching tacticals, such as {\tt
+REPEAT}, have built-in tracing options.
+\begin{description}
+\item[\ttindexbold{pause_tac}] 
+prints {\tt** Press RETURN to continue:} and then reads a line from the
+terminal.  If this line is blank then it returns the proof state unchanged;
+otherwise it fails (which may terminate a repetition).
+
+\item[\ttindexbold{print_tac}] 
+returns the proof state unchanged, with the side effect of printing it at
+the terminal.
+\end{description}
+
+
+\subsection{Sequences}
+\index{sequences (lazy lists)|bold}
+The module \ttindex{Sequence} declares a type of lazy lists.  It uses
+Isabelle's type \ttindexbold{option} to represent the possible presence
+(\ttindexbold{Some}) or absence (\ttindexbold{None}) of
+a value:
+\begin{ttbox}
+datatype 'a option = None  |  Some of 'a;
+\end{ttbox}
+
+\subsubsection{Basic operations on sequences}
+\begin{ttbox} 
+Sequence.null   : 'a Sequence.seq
+Sequence.seqof  : (unit -> ('a * 'a Sequence.seq) option)
+                     -> 'a Sequence.seq
+Sequence.single : 'a -> 'a Sequence.seq
+Sequence.pull   : 'a Sequence.seq -> ('a * 'a Sequence.seq) option
+\end{ttbox}
+\begin{description}
+\item[{\tt Sequence.null}] 
+is the empty sequence.
+
+\item[\tt Sequence.seqof (fn()=> Some($x$,$s$))] 
+constructs the sequence with head~$x$ and tail~$s$, neither of which is
+evaluated.
+
+\item[{\tt Sequence.single} $x$] 
+constructs the sequence containing the single element~$x$.
+
+\item[{\tt Sequence.pull} $s$] 
+returns {\tt None} if the sequence is empty and {\tt Some($x$,$s'$)} if the
+sequence has head~$x$ and tail~$s'$.  Warning: calling \hbox{Sequence.pull
+$s$} again will {\bf recompute} the value of~$x$; it is not stored!
+\end{description}
+
+
+\subsubsection{Converting between sequences and lists}
+\begin{ttbox} 
+Sequence.chop: int * 'a Sequence.seq -> 'a list * 'a Sequence.seq
+Sequence.list_of_s : 'a Sequence.seq -> 'a list
+Sequence.s_of_list : 'a list -> 'a Sequence.seq
+\end{ttbox}
+\begin{description}
+\item[{\tt Sequence.chop} ($n$, $s$)] 
+returns the first~$n$ elements of~$s$ as a list, paired with the remaining
+elements of~$s$.  If $s$ has fewer than~$n$ elements, then so will the
+list.
+
+\item[{\tt Sequence.list_of_s} $s$] 
+returns the elements of~$s$, which must be finite, as a list.
+
+\item[{\tt Sequence.s_of_list} $l$] 
+creates a sequence containing the elements of~$l$.
+\end{description}
+
+
+\subsubsection{Combining sequences}
+\begin{ttbox} 
+Sequence.append: 'a Sequence.seq * 'a Sequence.seq -> 'a Sequence.seq
+Sequence.interleave : 'a Sequence.seq * 'a Sequence.seq
+                                                   -> 'a Sequence.seq
+Sequence.flats   : 'a Sequence.seq Sequence.seq -> 'a Sequence.seq
+Sequence.maps    : ('a -> 'b) -> 'a Sequence.seq -> 'b Sequence.seq
+Sequence.filters : ('a -> bool) -> 'a Sequence.seq -> 'a Sequence.seq
+\end{ttbox} 
+\begin{description}
+\item[{\tt Sequence.append} ($s@1$, $s@2$)] 
+concatenates $s@1$ to $s@2$.
+
+\item[{\tt Sequence.interleave} ($s@1$, $s@2$)] 
+joins $s@1$ with $s@2$ by interleaving their elements.  The result contains
+all the elements of the sequences, even if both are infinite.
+
+\item[{\tt Sequence.flats} $ss$] 
+concatenates a sequence of sequences.
+
+\item[{\tt Sequence.maps} $f$ $s$] 
+applies $f$ to every element of~$s=x@1,x@2,\ldots$, yielding the sequence
+$f(x@1),f(x@2),\ldots$.
+
+\item[{\tt Sequence.filters} $p$ $s$] 
+returns the sequence consisting of all elements~$x$ of~$s$ such that $p(x)$
+is {\tt true}.
+\end{description}
+
+\index{tactics|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/tctical.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,505 @@
+%% $Id$
+\chapter{Tacticals}
+\index{tacticals|(}
+\index{tactics!combining|see{tacticals}}
+Tacticals are operations on tactics.  Their implementation makes use of
+functional programming techniques, especially for sequences.  Most of the
+time, you may forget about this and regard tacticals as high-level control
+structures.
+
+\section{The basic tacticals}
+\subsection{Joining two tactics}
+\index{tacticals!joining two tactics|bold}
+The tacticals {\tt THEN} and {\tt ORELSE}, which provide sequencing and
+alternation, underlie most of the other control structures in Isabelle.
+{\tt APPEND} and {\tt INTLEAVE} provide more sophisticated forms of
+alternation.
+\begin{ttbox} 
+THEN     : tactic * tactic -> tactic                 \hfill{\bf infix 1}
+ORELSE   : tactic * tactic -> tactic                 \hfill{\bf infix}
+APPEND   : tactic * tactic -> tactic                 \hfill{\bf infix}
+INTLEAVE : tactic * tactic -> tactic                 \hfill{\bf infix}
+\end{ttbox}
+\begin{description}
+\item[\tt $tac@1$ \ttindexbold{THEN} $tac@2$] 
+is the sequential composition of the two tactics.  Applied to a proof
+state, it returns all states reachable in two steps by applying $tac@1$
+followed by~$tac@2$.  First, it applies $tac@1$ to the proof state, getting a
+sequence of next states; then, it applies $tac@2$ to each of these and
+concatenates the results.
+
+\item[\tt $tac@1$ \ttindexbold{ORELSE} $tac@2$] 
+makes a choice between the two tactics.  Applied to a state, it
+tries~$tac@1$ and returns the result if non-empty; if $tac@1$ fails then it
+uses~$tac@2$.  This is a deterministic choice: if $tac@1$ succeeds then
+$tac@2$ is excluded.
+
+\item[\tt $tac@1$ \ttindexbold{APPEND} $tac@2$] 
+concatenates the results of $tac@1$ and~$tac@2$.  By not making a commitment
+to either tactic, {\tt APPEND} helps avoid incompleteness during search.
+
+\item[\tt $tac@1$ \ttindexbold{INTLEAVE} $tac@2$] 
+interleaves the results of $tac@1$ and~$tac@2$.  Thus, it includes all
+possible next states, even if one of the tactics returns an infinite
+sequence.
+\end{description}
+
+
+\subsection{Joining a list of tactics}
+\index{tacticals!joining a list of tactics|bold}
+\begin{ttbox} 
+EVERY : tactic list -> tactic
+FIRST : tactic list -> tactic
+\end{ttbox}
+{\tt EVERY} and {\tt FIRST} are block structured versions of {\tt THEN} and
+{\tt ORELSE}\@.
+\begin{description}
+\item[\ttindexbold{EVERY} {$[tac@1,\ldots,tac@n]$}] 
+abbreviates \hbox{\tt$tac@1$ THEN \ldots{} THEN $tac@n$}.  It is useful for
+writing a series of tactics to be executed in sequence.
+
+\item[\ttindexbold{FIRST} {$[tac@1,\ldots,tac@n]$}] 
+abbreviates \hbox{\tt$tac@1$ ORELSE \ldots{} ORELSE $tac@n$}.  It is useful for
+writing a series of tactics to be attempted one after another.
+\end{description}
+
+
+\subsection{Repetition tacticals}
+\index{tacticals!repetition|bold}
+\begin{ttbox} 
+TRY           : tactic -> tactic
+REPEAT_DETERM : tactic -> tactic
+REPEAT        : tactic -> tactic
+REPEAT1       : tactic -> tactic
+trace_REPEAT  : bool ref \hfill{\bf initially false}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{TRY} {\it tac}] 
+applies {\it tac\/} to the proof state and returns the resulting sequence,
+if non-empty; otherwise it returns the original state.  Thus, it applies
+{\it tac\/} at most once.
+
+\item[\ttindexbold{REPEAT_DETERM} {\it tac}] 
+applies {\it tac\/} to the proof state and, recursively, to the head of the
+resulting sequence.  It returns the first state to make {\it tac\/} fail.
+It is deterministic, discarding alternative outcomes.
+
+\item[\ttindexbold{REPEAT} {\it tac}] 
+applies {\it tac\/} to the proof state and, recursively, to each element of
+the resulting sequence.  The resulting sequence consists of those states
+that make {\it tac\/} fail.  Thus, it applies {\it tac\/} as many times as
+possible (including zero times), and allows backtracking over each
+invocation of {\it tac}.  It is more general than {\tt REPEAT_DETERM}, but
+requires more space.
+
+\item[\ttindexbold{REPEAT1} {\it tac}] 
+is like \hbox{\tt REPEAT {\it tac}} but it always applies {\it tac\/} at
+least once, failing if this is impossible.
+
+\item[\ttindexbold{trace_REPEAT} \tt:= true;] 
+enables an interactive tracing mode for {\tt REPEAT_DETERM} and {\tt
+REPEAT}.  To view the tracing options, type {\tt h} at the prompt.
+\end{description}
+
+
+\subsection{Identities for tacticals}
+\index{tacticals!identities for|bold}
+\begin{ttbox} 
+all_tac : tactic
+no_tac  : tactic
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{all_tac}] 
+maps any proof state to the one-element sequence containing that state.
+Thus, it succeeds for all states.  It is the identity element of the
+tactical \ttindex{THEN}\@.
+
+\item[\ttindexbold{no_tac}] 
+maps any proof state to the empty sequence.  Thus it succeeds for no state.
+It is the identity element of \ttindex{ORELSE}, \ttindex{APPEND}, and 
+\ttindex{INTLEAVE}\@.  Also, it is a zero element for \ttindex{THEN}, which means that
+\hbox{\tt$tac$ THEN no_tac} is equivalent to {\tt no_tac}.
+\end{description}
+These primitive tactics are useful when writing tacticals.  For example,
+\ttindexbold{TRY} and \ttindexbold{REPEAT} (ignoring tracing) can be coded
+as follows: 
+\begin{ttbox} 
+fun TRY tac = tac ORELSE all_tac;
+
+fun REPEAT tac = Tactic
+     (fn state => tapply((tac THEN REPEAT tac) ORELSE all_tac,
+                         state));
+\end{ttbox}
+If $tac$ can return multiple outcomes then so can \hbox{\tt REPEAT $tac$}.
+Since {\tt REPEAT} uses \ttindex{ORELSE} and not {\tt APPEND} or {\tt
+INTLEAVE}, it applies $tac$ as many times as possible in each
+outcome.
+
+\begin{warn}
+Note {\tt REPEAT}'s explicit abstraction over the proof state.  Recursive
+tacticals must be coded in this awkward fashion to avoid infinite
+recursion.  With the following definition, \hbox{\tt REPEAT $tac$} would
+loop:
+\begin{ttbox} 
+fun REPEAT tac = (tac THEN REPEAT tac) ORELSE all_tac;
+\end{ttbox}
+\par\noindent
+The built-in {\tt REPEAT} avoids~{\tt THEN}, handling sequences explicitly
+and using tail recursion.  This sacrifices clarity, but saves much space by
+discarding intermediate proof states.
+\end{warn}
+
+
+\section{Control and search tacticals}
+A predicate on theorems, namely a function of type \hbox{\tt thm->bool},
+can test whether a proof state enjoys some desirable property --- such as
+having no subgoals.  Tactics that search for satisfactory states are easy
+to express.  The main search procedures, depth-first, breadth-first and
+best-first, are provided as tacticals.  They generate the search tree by
+repeatedly applying a given tactic.
+
+
+\subsection{Filtering a tactic's results}
+\index{tacticals!for filtering|bold}
+\index{tactics!filtering results of|bold}
+\begin{ttbox} 
+FILTER  : (thm -> bool) -> tactic -> tactic
+CHANGED : tactic -> tactic
+\end{ttbox}
+\begin{description}
+\item[\tt \tt FILTER {\it p} $tac$] 
+applies $tac$ to the proof state and returns a sequence consisting of those
+result states that satisfy~$p$.
+
+\item[\ttindexbold{CHANGED} {\it tac}] 
+applies {\it tac\/} to the proof state and returns precisely those states
+that differ from the original state.  Thus, \hbox{\tt CHANGED {\it tac}}
+always has some effect on the state.
+\end{description}
+
+
+\subsection{Depth-first search}
+\index{tacticals!searching|bold}
+\index{tracing!of searching tacticals}
+\begin{ttbox} 
+DEPTH_FIRST   : (thm->bool) -> tactic -> tactic
+DEPTH_SOLVE   : tactic -> tactic
+DEPTH_SOLVE_1 : tactic -> tactic
+trace_DEPTH_FIRST: bool ref \hfill{\bf initially false}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{DEPTH_FIRST} {\it satp} {\it tac}] 
+returns the proof state if {\it satp} returns true.  Otherwise it applies
+{\it tac}, then recursively searches from each element of the resulting
+sequence.  The code uses a stack for efficiency, in effect applying
+\hbox{\tt {\it tac} THEN DEPTH_FIRST {\it satp} {\it tac}} to the state.
+
+\item[\ttindexbold{DEPTH_SOLVE} {\it tac}] 
+uses {\tt DEPTH_FIRST} to search for states having no subgoals.
+
+\item[\ttindexbold{DEPTH_SOLVE_1} {\it tac}] 
+uses {\tt DEPTH_FIRST} to search for states having fewer subgoals than the
+given state.  Thus, it insists upon solving at least one subgoal.
+
+\item[\ttindexbold{trace_DEPTH_FIRST} \tt:= true;] 
+enables interactive tracing for {\tt DEPTH_FIRST}.  To view the
+tracing options, type {\tt h} at the prompt.
+\end{description}
+
+
+\subsection{Other search strategies}
+\index{tacticals!searching|bold}
+\index{tracing!of searching tacticals}
+\begin{ttbox} 
+BREADTH_FIRST   : (thm->bool) -> tactic -> tactic
+BEST_FIRST      : (thm->bool)*(thm->int) -> tactic -> tactic
+THEN_BEST_FIRST : tactic * ((thm->bool) * (thm->int) * tactic)
+                  -> tactic                    \hfill{\bf infix 1}
+trace_BEST_FIRST: bool ref \hfill{\bf initially false}
+\end{ttbox}
+These search strategies will find a solution if one exists.  However, they
+do not enumerate all solutions; they terminate after the first satisfactory
+result from {\it tac}.
+\begin{description}
+\item[\ttindexbold{BREADTH_FIRST} {\it satp} {\it tac}] 
+uses breadth-first search to find states for which {\it satp\/} is true.
+For most applications, it is too slow.
+
+\item[\ttindexbold{BEST_FIRST} $(satp,distf)$ {\it tac}] 
+does a heuristic search, using {\it distf\/} to estimate the distance from
+a satisfactory state.  It maintains a list of states ordered by distance.
+It applies $tac$ to the head of this list; if the result contains any
+satisfactory states, then it returns them.  Otherwise, {\tt BEST_FIRST}
+adds the new states to the list, and continues.  
+
+The distance function is typically \ttindex{size_of_thm}, which computes
+the size of the state.  The smaller the state, the fewer and simpler
+subgoals it has.
+
+\item[$tac@0$ \ttindexbold{THEN_BEST_FIRST} $(satp,distf,tac)$] 
+is like {\tt BEST_FIRST}, except that the priority queue initially
+contains the result of applying $tac@0$ to the proof state.  This tactical
+permits separate tactics for starting the search and continuing the search.
+
+\item[\ttindexbold{trace_BEST_FIRST} \tt:= true;] 
+enables an interactive tracing mode for {\tt BEST_FIRST}.  To view the
+tracing options, type {\tt h} at the prompt.
+\end{description}
+
+
+\subsection{Auxiliary tacticals for searching}
+\index{tacticals!conditional}
+\index{tacticals!deterministic}
+\begin{ttbox} 
+COND        : (thm->bool) -> tactic -> tactic -> tactic
+IF_UNSOLVED : tactic -> tactic
+DETERM      : tactic -> tactic
+\end{ttbox}
+\begin{description}
+\item[\tt \tt COND {\it p} $tac@1$ $tac@2$] 
+applies $tac@1$ to the proof state if it satisfies~$p$, and applies $tac@2$
+otherwise.  It is a conditional tactical in that only one of $tac@1$ and
+$tac@2$ is applied to a proof state.  However, both $tac@1$ and $tac@2$ are
+evaluated because \ML{} uses eager evaluation.
+
+\item[\ttindexbold{IF_UNSOLVED} {\it tac}] 
+applies {\it tac\/} to the proof state if it has any subgoals, and simply
+returns the proof state otherwise.  Many common tactics, such as {\tt
+resolve_tac}, fail if applied to a proof state that has no subgoals.
+
+\item[\ttindexbold{DETERM} {\it tac}] 
+applies {\it tac\/} to the proof state and returns the head of the
+resulting sequence.  {\tt DETERM} limits the search space by making its
+argument deterministic.
+\end{description}
+
+
+\subsection{Predicates and functions useful for searching}
+\index{theorems!size of}
+\index{theorems!equality of}
+\begin{ttbox} 
+has_fewer_prems : int -> thm -> bool
+eq_thm          : thm * thm -> bool
+size_of_thm     : thm -> int
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{has_fewer_prems} $n$ $thm$] 
+reports whether $thm$ has fewer than~$n$ premises.  By currying,
+\hbox{\tt has_fewer_prems $n$} is a predicate on theorems; it may 
+be given to the searching tacticals.
+
+\item[\ttindexbold{eq_thm}($thm1$,$thm2$)] 
+reports whether $thm1$ and $thm2$ are equal.  Both theorems must have
+identical signatures.  Both theorems must have the same conclusions, and
+the same hypotheses, in the same order.  Names of bound variables are
+ignored.
+
+\item[\ttindexbold{size_of_thm} $thm$] 
+computes the size of $thm$, namely the number of variables, constants and
+abstractions in its conclusion.  It may serve as a distance function for 
+\ttindex{BEST_FIRST}. 
+\end{description}
+
+
+\section{Tacticals for subgoal numbering}
+When conducting a backward proof, we normally consider one goal at a time.
+A tactic can affect the entire proof state, but many tactics --- such as
+{\tt resolve_tac} and {\tt assume_tac} --- work on a single subgoal.
+Subgoals are designated by a positive integer, so Isabelle provides
+tacticals for combining values of type {\tt int->tactic}.
+
+
+\subsection{Restricting a tactic to one subgoal}
+\index{tactics!restricting to a subgoal}
+\index{tacticals!for restriction to a subgoal}
+\begin{ttbox} 
+SELECT_GOAL : tactic -> int -> tactic
+METAHYPS    : (thm list -> tactic) -> int -> tactic
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{SELECT_GOAL} {\it tac} $i$] 
+restricts the effect of {\it tac\/} to subgoal~$i$ of the proof state.  It
+fails if there is no subgoal~$i$, or if {\it tac\/} changes the main goal
+(do not use {\tt rewrite_tac}).  It applies {\it tac\/} to a dummy proof
+state and uses the result to refine the original proof state at
+subgoal~$i$.  If {\it tac\/} returns multiple results then so does 
+\hbox{\tt SELECT_GOAL {\it tac} $i$}.
+
+Its workings are devious.  {\tt SELECT_GOAL} creates a state of the
+form $\phi\Imp\phi$, with one subgoal.  If subgoal~$i$ has the form say
+$\psi\Imp\theta$, then {\tt SELECT_GOAL} creates the state 
+\[ (\Forall x.\psi\Imp\theta)\Imp(\Forall x.\psi\Imp\theta) \]
+rather than $\List{\psi\Imp\theta;\; \psi}\Imp\theta$, which has two
+subgoals.
+
+\item[\ttindexbold{METAHYPS} {\it tacf} $i$] 
+takes subgoal~$i$, of the form 
+\[ \Forall x@1 \ldots x@l. \List{\theta@1; \ldots; \theta@k}\Imp\theta, \]
+and creates the list $\theta'@1$, \ldots, $\theta'@k$ of meta-level
+assumptions.  In these theorems, the subgoal's parameters ($x@1$,
+\ldots,~$x@l$) become free variables.  It supplies the assumptions to
+$tacf$ and applies the resulting tactic to the proof state
+$\theta\Imp\theta$.
+
+If the resulting proof state is $\List{\phi@1; \ldots; \phi@n} \Imp \phi$,
+possibly containing $\theta'@1,\ldots,\theta'@k$ as assumptions, then it is
+lifted back into the original context, yielding $n$ subgoals.
+
+Meta-level assumptions may not contain unknowns.  Unknowns in
+$\theta@1,\ldots,\theta@k$ become free variables in $\theta'@1$, \ldots,
+$\theta'@k$, and are restored afterwards; the {\tt METAHYPS} call cannot
+instantiate them.  Unknowns in $\theta$ may be instantiated.  New unknowns
+in $\phi@1$, \ldots; $\phi@n$ are lifted over the parameters.
+
+Here is a typical application.  Calling {\tt hyp_res_tac}~$i$ resolves
+subgoal~$i$ with one of its own assumptions, which may itself have the form
+of an inference rule (these are called {\bf higher-level assumptions}).  
+\begin{ttbox} 
+val hyp_res_tac = METAHYPS (fn prems => resolve_tac prems 1);
+\end{ttbox} 
+\end{description}
+
+\begin{warn}
+{\tt METAHYPS} fails if the context or new subgoals contain type unknowns.
+In principle, the tactical could treat these like ordinary unknowns.
+\end{warn}
+
+
+\subsection{Scanning for a subgoal by number}
+\index{tacticals!scanning for subgoals|bold}
+\begin{ttbox} 
+ALLGOALS         : (int -> tactic) -> tactic
+TRYALL           : (int -> tactic) -> tactic
+SOMEGOAL         : (int -> tactic) -> tactic
+FIRSTGOAL        : (int -> tactic) -> tactic
+REPEAT_SOME      : (int -> tactic) -> tactic
+REPEAT_FIRST     : (int -> tactic) -> tactic
+trace_goalno_tac : (int -> tactic) -> int -> tactic
+\end{ttbox}
+These apply a tactic function of type {\tt int -> tactic} to all the
+subgoal numbers of a proof state, and join the resulting tactics using
+\ttindex{THEN} or \ttindex{ORELSE}\@.  Thus, they apply the tactic to all the
+subgoals, or to one subgoal.  
+
+Suppose that the original proof state has $n$ subgoals.
+
+\begin{description}
+\item[\ttindexbold{ALLGOALS} {\it tacf}] 
+is equivalent to
+\hbox{\tt$tacf(n)$ THEN \ldots{} THEN $tacf(1)$}.  
+
+It applies {\it tacf} to all the subgoals, counting {\bf downwards} (to
+avoid problems when subgoals are added or deleted).
+
+\item[\ttindexbold{TRYALL} {\it tacf}] 
+is equivalent to
+\hbox{\tt TRY($tacf(n)$) THEN \ldots{} THEN TRY($tacf(1)$)}.  
+
+It attempts to apply {\it tacf} to all the subgoals.  For instance,
+\hbox{\tt TRYALL assume_tac} attempts to solve all the subgoals by
+assumption.
+
+\item[\ttindexbold{SOMEGOAL} {\it tacf}] 
+is equivalent to
+\hbox{\tt$tacf(n)$ ORELSE \ldots{} ORELSE $tacf(1)$}.  
+
+It applies {\it tacf} to one subgoal, counting {\bf downwards}.  
+\hbox{\tt SOMEGOAL assume_tac} solves one subgoal by assumption, failing if
+this is impossible.
+
+\item[\ttindexbold{FIRSTGOAL} {\it tacf}] 
+is equivalent to
+\hbox{\tt$tacf(1)$ ORELSE \ldots{} ORELSE $tacf(n)$}.  
+
+It applies {\it tacf} to one subgoal, counting {\bf upwards}.
+
+\item[\ttindexbold{REPEAT_SOME} {\it tacf}] 
+applies {\it tacf} once or more to a subgoal, counting {\bf downwards}.
+
+\item[\ttindexbold{REPEAT_FIRST} {\it tacf}] 
+applies {\it tacf} once or more to a subgoal, counting {\bf upwards}.
+
+\item[\ttindexbold{trace_goalno_tac} {\it tac} {\it i}] 
+applies \hbox{\it tac i\/} to the proof state.  If the resulting sequence
+is non-empty, then it is returned, with the side-effect of printing {\tt
+Subgoal~$i$ selected}.  Otherwise, {\tt trace_goalno_tac} returns the empty
+sequence and prints nothing.
+
+It indicates that ``the tactic worked for subgoal~$i$'' and is mainly used
+with {\tt SOMEGOAL} and {\tt FIRSTGOAL}.
+\end{description}
+
+
+\subsection{Joining tactic functions}
+\index{tacticals!joining tactic functions|bold}
+\begin{ttbox} 
+THEN'     : ('a -> tactic) * ('a -> tactic) -> 'a -> tactic \hfill{\bf infix 1}
+ORELSE'   : ('a -> tactic) * ('a -> tactic) -> 'a -> tactic \hfill{\bf infix}
+APPEND'   : ('a -> tactic) * ('a -> tactic) -> 'a -> tactic \hfill{\bf infix}
+INTLEAVE' : ('a -> tactic) * ('a -> tactic) -> 'a -> tactic \hfill{\bf infix}
+EVERY'    : ('a -> tactic) list -> 'a -> tactic
+FIRST'    : ('a -> tactic) list -> 'a -> tactic
+\end{ttbox}
+These help to express tactics that specify subgoal numbers.  The tactic
+\begin{ttbox} 
+SOMEGOAL (fn i => resolve_tac rls i  ORELSE  eresolve_tac erls i)
+\end{ttbox}
+can be simplified to
+\begin{ttbox} 
+SOMEGOAL (resolve_tac rls  ORELSE'  eresolve_tac erls)
+\end{ttbox}
+Note that {\tt TRY'}, {\tt REPEAT'}, {\tt DEPTH_FIRST'}, etc.\ are not
+provided, because function composition accomplishes the same purpose.
+The tactic
+\begin{ttbox} 
+ALLGOALS (fn i => REPEAT (etac exE i  ORELSE  atac i))
+\end{ttbox}
+can be simplified to
+\begin{ttbox} 
+ALLGOALS (REPEAT o (etac exE  ORELSE'  atac))
+\end{ttbox}
+These tacticals are polymorphic; $x$ need not be an integer.
+\begin{center} \tt
+\begin{tabular}{r@{\rm\ \ yields\ \ }l}
+    ($tacf@1$~~THEN'~~$tacf@2$)$(x)$ \index{*THEN'} &
+    $tacf@1(x)$~~THEN~~$tacf@2(x)$ \\
+
+    ($tacf@1$ ORELSE' $tacf@2$)$(x)$ \index{*ORELSE'} &
+    $tacf@1(x)$ ORELSE $tacf@2(x)$ \\
+
+    ($tacf@1$ APPEND' $tacf@2$)$(x)$ \index{*APPEND'} &
+    $tacf@1(x)$ APPEND $tacf@2(x)$ \\
+
+    ($tacf@1$ INTLEAVE' $tacf@2$)$(x)$ \index{*INTLEAVE'} &
+    $tacf@1(x)$ INTLEAVE $tacf@2(x)$ \\
+
+    EVERY' $[tacf@1,\ldots,tacf@n] \; (x)$ \index{*EVERY'} &
+    EVERY $[tacf@1(x),\ldots,tacf@n(x)]$ \\
+
+    FIRST' $[tacf@1,\ldots,tacf@n] \; (x)$ \index{*FIRST'} &
+    FIRST $[tacf@1(x),\ldots,tacf@n(x)]$
+\end{tabular}
+\end{center}
+
+
+\subsection{Applying a list of tactics to 1}
+\index{tacticals!joining tactic functions|bold}
+\begin{ttbox} 
+EVERY1: (int -> tactic) list -> tactic
+FIRST1: (int -> tactic) list -> tactic
+\end{ttbox}
+A common proof style is to treat the subgoals as a stack, always
+restricting attention to the first subgoal.  Such proofs contain long lists
+of tactics, each applied to~1.  These can be simplified using {\tt EVERY1}
+and {\tt FIRST1}:
+\begin{center} \tt
+\begin{tabular}{r@{\rm\ \ abbreviates\ \ }l}
+    EVERY1 $[tacf@1,\ldots,tacf@n]$ \indexbold{*EVERY1} &
+    EVERY $[tacf@1(1),\ldots,tacf@n(1)]$ \\
+
+    FIRST1 $[tacf@1,\ldots,tacf@n]$ \indexbold{*FIRST1} &
+    FIRST $[tacf@1(1),\ldots,tacf@n(1)]$
+\end{tabular}
+\end{center}
+
+\index{tacticals|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/theories.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,445 @@
+%% $Id$
+\chapter{Theories, Terms and Types} \label{theories}
+\index{theories|(}\index{signatures|bold}
+\index{reading!axioms|see{{\tt extend_theory} and {\tt assume_ax}}}
+Theories organize the syntax, declarations and axioms of a mathematical
+development.  They are built, starting from the Pure theory, by extending
+and merging existing theories.  They have the \ML{} type \ttindex{theory}.
+Theory operations signal errors by raising exception \ttindex{THEORY},
+returning a message and a list of theories.
+
+Signatures, which contain information about sorts, types, constants and
+syntax, have the \ML{} type~\ttindexbold{Sign.sg}.  For identification,
+each signature carries a unique list of {\bf stamps}, which are~\ML{}
+references (to strings).  The strings serve as human-readable names; the
+references serve as unique identifiers.  Each primitive signature has a
+single stamp.  When two signatures are merged, their lists of stamps are
+also merged.  Every theory carries a unique signature.
+
+Terms and types are the underlying representation of logical syntax.  Their
+\ML{} definitions are irrelevant to naive Isabelle users.  Programmers who wish
+to extend Isabelle may need to know such details, say to code a tactic that
+looks for subgoals of a particular form.  Terms and types may be
+`certified' to be well-formed with respect to a given signature.
+
+\section{Defining Theories}
+\label{DefiningTheories}
+
+Theories can be defined either using concrete syntax or by calling certain
+\ML-functions (see \S\ref{BuildingATheory}).  Figure~\ref{TheorySyntax}
+presents the concrete syntax for theories.  This grammar employs the
+following conventions: 
+\begin{itemize}
+\item Identifiers such as $theoryDef$ denote nonterminal symbols.
+\item Text in {\tt typewriter font} denotes terminal symbols.
+\item \ldots{} indicates repetition of a phrase.
+\item A vertical bar~($|$) separates alternative phrases.
+\item Square [brackets] enclose optional phrases.
+\item $id$ stands for an Isabelle identifier.
+\item $string$ stands for an ML string, with its quotation marks.
+\item $k$ stands for a natural number.
+\item $text$ stands for arbitrary ML text.
+\end{itemize}
+
+\begin{figure}[hp]
+\begin{center}
+\begin{tabular}{rclc}
+$theoryDef$ &=& $id$ {\tt=} $id@1$ {\tt+} \dots {\tt+} $id@n$
+                            [{\tt+} $extension$]\\\\
+$extension$ &=& [$classes$] [$default$] [$types$] [$arities$] [$consts$]
+                [$rules$] {\tt end} [$ml$]\\\\
+$classes$ &=& \ttindex{classes} $class$ \dots $class$ \\\\
+$class$ &=& $id$ [{\tt<} $id@1${\tt,} \dots{\tt,} $id@n$] \\\\
+$default$ &=& \ttindex{default} $sort$ \\\\
+$sort$ &=& $id$ ~~$|$~~ {\tt\{} $id@1${\tt,} \dots{\tt,} $id@n$ {\tt\}} \\\\
+$name$ &=& $id$ ~~$|$~~ $string$ \\\\
+$types$ &=& \ttindex{types} $typeDecl$ \dots $typeDecl$ \\\\
+$typeDecl$ &=& $name${\tt,} \dots{\tt,} $name$ $k$
+               [{\tt(} $infix$ {\tt)}] \\\\
+$infix$ &=& \ttindex{infixl} $k$ ~~$|$~~ \ttindex{infixr} $k$ \\\\
+$arities$ &=& \ttindex{arities} $arityDecl$ \dots $arityDecl$ \\\\
+$arityDecl$ &=& $name${\tt,} \dots{\tt,} $name$ {\tt::} $arity$ \\\\
+$arity$ &=& [{\tt(} $sort${\tt,} \dots{\tt,} $sort$ {\tt)}] $id$ \\\\
+$consts$ &=& \ttindex{consts} $constDecl$ \dots $constDecl$ \\\\
+$constDecl$ &=& $name@1${\tt,} \dots{\tt,} $name@n$ {\tt::} $string$
+                [{\tt(} $mixfix$ {\tt)}] \\\\
+$mixfix$ &=& $string$
+             [ [\quad{\tt[} $k@1${\tt,} \dots{\tt,} $k@n$ {\tt]}\quad] $k$]\\
+       &$|$& $infix$ \\
+       &$|$& \ttindex{binder} $string$ $k$\\\\
+$rules$ &=& \ttindex{rules} $rule$ \dots $rule$ \\\\
+$rule$ &=& $id$ $string$ \\\\
+$ml$ &=& \ttindex{ML} $text$
+\end{tabular}
+\end{center}
+\caption{Theory Syntax}
+\label{TheorySyntax}
+\end{figure}
+
+The different parts of a theory definition are interpreted as follows:
+\begin{description}
+\item[$theoryDef$] A theory has a name $id$ and is an extension of the union
+  of existing theories $id@1$ \dots $id@n$ with new classes, types,
+  constants, syntax and axioms.  The basic theory, which contains only the
+  meta-logic, is called {\tt Pure}.
+\item[$class$] The new class $id$ is declared as a subclass of existing
+  classes $id@1$ \dots $id@n$.  This rules out cyclic class structures.
+  Isabelle automatically computes the transitive closure of subclass
+  hierarchies.  Hence it is not necessary to declare $c@1 < c@3$ in addition
+  to $c@1 < c@2$ and $c@2 < c@3$.
+\item[$default$] introduces $sort$ as the new default sort for type
+  variables.  Unconstrained type variables in an input string are
+  automatically constrained by $sort$; this does not apply to type variables
+  created internally during type inference.  If omitted,
+  the default sort is the same as in the parent theory.
+\item[$sort$] is a finite set $id@1$ \dots $id@n$ of classes; a single class
+  $id$ abbreviates the singleton set {\tt\{}$id${\tt\}}.
+\item[$name$] is either an alphanumeric identifier or an arbitrary string.
+\item[$typeDecl$] Each $name$ is declared as a new type constructor with
+  $k$ arguments.  Only binary type constructors can have infix status and
+  symbolic names ($string$).
+\item[$infix$] declares a type or constant to be an infix operator of
+  precedence $k$ associating to the left ({\tt infixl}) or right ({\tt
+    infixr}).
+\item[$arityDecl$] Each existing type constructor $name@1$ \dots $name@n$
+  is given the additional arity $arity$.
+\item[$constDecl$] The new constants $name@1$ \dots $name@n$ are declared to
+  be of type $string$.  For display purposes they can be annotated with
+  $mixfix$ declarations.
+\item[$mixfix$] $string$ is a mixfix template of the form {\tt"}\dots{\tt\_}
+  \dots{\tt\_} \dots{\tt"} where the $i$-th underscore indicates the position
+  where the $i$-th argument should go, $k@i$ is the minimum precedence of
+  the $i$-th argument, and $k$ the precedence of the construct.  The list
+  \hbox{\tt[$k@1$, \dots, $k@n$]} is optional.
+
+  Binary constants can be given infix status.
+
+  A constant $f$ {\tt::} $(\tau@1\To\tau@2)\To\tau@3$ can be given {\em
+    binder} status: {\tt binder} $Q$ $p$ causes $Q\,x.F(x)$ to be treated
+  like $f(F)$; $p$ is the precedence of the construct.
+\item[$rule$] A rule consists of a name $id$ and a formula $string$.  Rule
+  names must be distinct.
+\item[$ml$] This final part of a theory consists of ML code, 
+  typically for parse and print translations.
+\end{description}
+The $mixfix$ and $ml$ sections are explained in more detail in the {\it
+Defining Logics} chapter of the {\it Logics} manual.
+
+\begin{ttbox} 
+use_thy: string -> unit
+\end{ttbox}
+Each theory definition must reside in a separate file.  Let the file {\it
+  tf}{\tt.thy} contain the definition of a theory called $TF$ with rules named
+$r@1$ \dots $r@n$.  Calling \ttindexbold{use_thy}~{\tt"}{\it tf\/}{\tt"} reads
+file {\it tf}{\tt.thy}, writes an intermediate {\ML}-file {\tt.}{\it
+  tf}{\tt.thy.ML}, and reads the latter file.  This declares an {\ML}
+structure~$TF$ containing a component {\tt thy} for the new theory
+and components $r@1$ \dots $r@n$ for the rules; it also contains the
+definitions of the {\tt ML} section if any.  Then {\tt use_thy}
+reads the file {\it tf}{\tt.ML}, if it exists; this file normally contains
+proofs involving the new theory.
+
+
+\section{Basic operations on theories}
+\subsection{Extracting an axiom from a theory}
+\index{theories!extracting axioms|bold}\index{axioms}
+\begin{ttbox} 
+get_axiom: theory -> string -> thm
+assume_ax: theory -> string -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{get_axiom} $thy$ $name$] 
+returns an axiom with the given $name$ from $thy$, raising exception
+\ttindex{THEORY} if none exists.  Merging theories can cause several axioms
+to have the same name; {\tt get_axiom} returns an arbitrary one.
+
+\item[\ttindexbold{assume_ax} $thy$ $formula$] 
+reads the {\it formula} using the syntax of $thy$, following the same
+conventions as axioms in a theory definition.  You can thus pretend that
+{\it formula} is an axiom; in fact, {\tt assume_ax} returns an assumption.
+You can use the resulting theorem like an axiom.  Note that 
+\ttindex{result} complains about additional assumptions, but 
+\ttindex{uresult} does not.
+
+For example, if {\it formula} is
+\hbox{\tt a=b ==> b=a} then the resulting theorem might have the form
+\hbox{\tt\frenchspacing ?a=?b ==> ?b=?a  [!!a b. a=b ==> b=a]}
+\end{description}
+
+\subsection{Building a theory}
+\label{BuildingATheory}
+\index{theories!constructing|bold}
+\begin{ttbox} 
+pure_thy: theory
+merge_theories: theory * theory -> theory
+extend_theory: theory -> string
+        -> (class * class list) list 
+           * sort
+           * (string list * int)list
+           * (string list * (sort list * class))list
+           * (string list * string)list * sext option
+        -> (string*string)list -> theory
+\end{ttbox}
+Type \ttindex{class} is a synonym for {\tt string}; type \ttindex{sort} is
+a synonym for \hbox{\tt class list}.
+\begin{description}
+\item[\ttindexbold{pure_thy}] contains just the types, constants, and syntax
+  of the meta-logic.  There are no axioms: meta-level inferences are carried
+  out by \ML\ functions.
+\item[\ttindexbold{merge_theories} ($thy@1$, $thy@2$)] merges the two
+  theories $thy@1$ and $thy@2$.  The resulting theory contains all types,
+  constants and axioms of the constituent theories; its default sort is the
+  union of the default sorts of the constituent theories.
+\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"}
+      ($classes$, $default$, $types$, $arities$, $consts$, $sextopt$) $rules$]
+\hfill\break   %%% include if line is just too short
+is the \ML-equivalent of the following theory definition:
+\begin{ttbox}
+\(T\) = \(thy\) +
+classes \(c\) < \(c@1\),\(\dots\),\(c@m\)
+        \dots
+default {\(d@1,\dots,d@r\)}
+types   \(tycon@1\),\dots,\(tycon@i\) \(n\)
+        \dots
+arities \(tycon@1'\),\dots,\(tycon@j'\) :: (\(s@1\),\dots,\(s@n\))\(c\)
+        \dots
+consts  \(b@1\),\dots,\(b@k\) :: \(\tau\)
+        \dots
+rules   \(name\) \(rule\)
+        \dots
+end
+\end{ttbox}
+where
+\begin{tabular}[t]{l@{~=~}l}
+$classes$ & \tt[("$c$",["$c@1$",\dots,"$c@m$"]),\dots] \\
+$default$ & \tt["$d@1$",\dots,"$d@r$"]\\
+$types$   & \tt[([$tycon@1$,\dots,$tycon@i$], $n$),\dots] \\
+$arities$ & \tt[([$tycon'@1$,\dots,$tycon'@j$], ([$s@1$,\dots,$s@n$],$c$)),\dots]
+\\
+$consts$  & \tt[([$b@1$,\dots,$b@k$],$\tau$),\dots] \\
+$rules$   & \tt[("$name$",$rule$),\dots]
+\end{tabular}
+
+If theories are defined as in \S\ref{DefiningTheories}, new syntax is added
+as mixfix annotations to constants.  Using {\tt extend_theory}, new syntax can
+be added via $sextopt$ which is either {\tt None}, or {\tt Some($sext$)}.  The
+latter case is not documented.
+
+$T$ identifies the theory internally.  When a theory is redeclared, say to
+change an incorrect axiom, bindings to the old axiom may persist.  Isabelle
+ensures that the old and new theories are not involved in the same proof.
+Attempting to combine different theories having the same name $T$ yields the
+fatal error
+\begin{center} \tt
+Attempt to merge different versions of theory: $T$
+\end{center}
+\end{description}
+
+
+\subsection{Inspecting a theory}
+\index{theories!inspecting|bold}
+\begin{ttbox} 
+print_theory  : theory -> unit
+axioms_of     : theory -> (string*thm)list
+parents_of    : theory -> theory list
+sign_of       : theory -> Sign.sg
+stamps_of_thy : theory -> string ref list
+\end{ttbox}
+These provide a simple means of viewing a theory's components.
+Unfortunately, there is no direct connection between a theorem and its
+theory.
+\begin{description}
+\item[\ttindexbold{print_theory} {\it thy}]  
+prints the theory {\it thy\/} at the terminal as a set of identifiers.
+
+\item[\ttindexbold{axioms_of} $thy$] 
+returns the axioms of~$thy$ and its ancestors.
+
+\item[\ttindexbold{parents_of} $thy$] 
+returns the parents of~$thy$.  This list contains zero, one or two
+elements, depending upon whether $thy$ is {\tt pure_thy}, 
+\hbox{\tt extend_theory $thy$} or \hbox{\tt merge_theories ($thy@1$, $thy@2$)}.
+
+\item[\ttindexbold{stamps_of_thy} $thy$]\index{signatures}
+returns the stamps of the signature associated with~$thy$.
+
+\item[\ttindexbold{sign_of} $thy$] 
+returns the signature associated with~$thy$.  It is useful with functions
+like {\tt read_instantiate_sg}, which take a signature as an argument.
+\end{description}
+
+
+\section{Terms}
+\index{terms|bold}
+Terms belong to the \ML{} type \ttindexbold{term}, which is a concrete datatype
+with six constructors: there are six kinds of term.
+\begin{ttbox}
+type indexname = string * int;
+infix 9 $;
+datatype term = Const of string * typ
+              | Free  of string * typ
+              | Var   of indexname * typ
+              | Bound of int
+              | Abs   of string * typ * term
+              | op $  of term * term;
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Const}($a$, $T$)] 
+is the {\bf constant} with name~$a$ and type~$T$.  Constants include
+connectives like $\land$ and $\forall$ (logical constants), as well as
+constants like~0 and~$Suc$.  Other constants may be required to define the
+concrete syntax of a logic.
+
+\item[\ttindexbold{Free}($a$, $T$)] 
+is the {\bf free variable} with name~$a$ and type~$T$.
+
+\item[\ttindexbold{Var}($v$, $T$)] 
+is the {\bf scheme variable} with indexname~$v$ and type~$T$.  An
+\ttindexbold{indexname} is a string paired with a non-negative index, or
+subscript; a term's scheme variables can be systematically renamed by
+incrementing their subscripts.  Scheme variables are essentially free
+variables, but may be instantiated during unification.
+
+\item[\ttindexbold{Bound} $i$] 
+is the {\bf bound variable} with de Bruijn index~$i$, which counts the
+number of lambdas, starting from zero, between a variable's occurrence and
+its binding.  The representation prevents capture of variables.  For more
+information see de Bruijn \cite{debruijn72} or
+Paulson~\cite[page~336]{paulson91}.
+
+\item[\ttindexbold{Abs}($a$, $T$, $u$)] 
+is the {\bf abstraction} with body~$u$, and whose bound variable has
+name~$a$ and type~$T$.  The name is used only for parsing and printing; it
+has no logical significance.
+
+\item[\tt $t$ \$ $u$] \index{$@{\tt\$}|bold}
+is the {\bf application} of~$t$ to~$u$.  
+\end{description}
+Application is written as an infix operator in order to aid readability.
+For example, here is an \ML{} pattern to recognize first-order formula of
+the form~$A\imp B$, binding the subformulae to~$A$ and~$B$:
+\begin{ttbox} 
+Const("Trueprop",_) $ (Const("op -->",_) $ A $ B)
+\end{ttbox}
+
+
+\section{Certified terms}
+\index{terms!certified|bold}\index{signatures}
+A term $t$ can be {\bf certified} under a signature to ensure that every
+type in~$t$ is declared in the signature and every constant in~$t$ is
+declared as a constant of the same type in the signature.  It must be
+well-typed and must not have any `loose' bound variable
+references.\footnote{This concerns the internal representation of variable
+binding using de Bruijn indexes.} Meta-rules such as {\tt forall_elim} take
+certified terms as arguments.
+
+Certified terms belong to the abstract type \ttindexbold{Sign.cterm}.
+Elements of the type can only be created through the certification process.
+In case of error, Isabelle raises exception~\ttindex{TERM}\@.
+
+\subsection{Printing terms}
+\index{printing!terms|bold}
+\begin{ttbox} 
+Sign.string_of_cterm :      Sign.cterm -> string
+Sign.string_of_term  : Sign.sg -> term -> string
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Sign.string_of_cterm} $ct$] 
+displays $ct$ as a string.
+
+\item[\ttindexbold{Sign.string_of_term} $sign$ $t$] 
+displays $t$ as a string, using the syntax of~$sign$.
+\end{description}
+
+\subsection{Making and inspecting certified terms}
+\begin{ttbox} 
+Sign.cterm_of   : Sign.sg -> term -> Sign.cterm
+Sign.read_cterm : Sign.sg -> string * typ -> Sign.cterm
+Sign.rep_cterm  : Sign.cterm -> \{T:typ, t:term, 
+                                 sign: Sign.sg, maxidx:int\}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Sign.cterm_of} $sign$ $t$] \index{signatures}
+certifies $t$ with respect to signature~$sign$.
+
+\item[\ttindexbold{Sign.read_cterm} $sign$ ($s$, $T$)] 
+reads the string~$s$ using the syntax of~$sign$, creating a certified term.
+The term is checked to have type~$T$; this type also tells the parser what
+kind of phrase to parse.
+
+\item[\ttindexbold{Sign.rep_cterm} $ct$] 
+decomposes $ct$ as a record containing its type, the term itself, its
+signature, and the maximum subscript of its unknowns.  The type and maximum
+subscript are computed during certification.
+\end{description}
+
+
+\section{Types}
+\index{types|bold}
+Types belong to the \ML{} type \ttindexbold{typ}, which is a concrete
+datatype with three constructors.  Types are classified by sorts, which are
+lists of classes.  A class is represented by a string.
+\begin{ttbox}
+type class = string;
+type sort  = class list;
+
+datatype typ = Type  of string * typ list
+             | TFree of string * sort
+             | TVar  of indexname * sort;
+
+infixr 5 -->;
+fun S --> T = Type("fun",[S,T]);
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Type}($a$, $Ts$)] 
+applies the {\bf type constructor} named~$a$ to the type operands~$Ts$.
+Type constructors include~\ttindexbold{fun}, the binary function space
+constructor, as well as nullary type constructors such
+as~\ttindexbold{prop}.  Other type constructors may be introduced.  In
+expressions, but not in patterns, \hbox{\tt$S$-->$T$} is a convenient
+shorthand for function types.
+
+\item[\ttindexbold{TFree}($a$, $s$)] 
+is the {\bf free type variable} with name~$a$ and sort~$s$.
+
+\item[\ttindexbold{TVar}($v$, $s$)] 
+is the {\bf scheme type variable} with indexname~$v$ and sort~$s$.  Scheme
+type variables are essentially free type variables, but may be instantiated
+during unification.
+\end{description}
+
+
+\section{Certified types}
+\index{types!certified|bold}
+Certified types, which are analogous to certified terms, have type 
+\ttindexbold{Sign.ctyp}.
+
+\subsection{Printing types}
+\index{printing!types|bold}
+\begin{ttbox} 
+Sign.string_of_ctyp :      Sign.ctyp -> string
+Sign.string_of_typ  : Sign.sg -> typ -> string
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Sign.string_of_ctyp} $cT$] 
+displays $cT$ as a string.
+
+\item[\ttindexbold{Sign.string_of_typ} $sign$ $T$] 
+displays $T$ as a string, using the syntax of~$sign$.
+\end{description}
+
+
+\subsection{Making and inspecting certified types}
+\begin{ttbox} 
+Sign.ctyp_of  : Sign.sg -> typ -> Sign.ctyp
+Sign.rep_ctyp : Sign.ctyp -> \{T: typ, sign: Sign.sg\}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{Sign.ctyp_of} $sign$ $T$] \index{signatures}
+certifies $T$ with respect to signature~$sign$.
+
+\item[\ttindexbold{Sign.rep_ctyp} $cT$] 
+decomposes $cT$ as a record containing the type itself and its signature.
+\end{description}
+
+\index{theories|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/thm.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,640 @@
+%% $Id$
+\chapter{Theorems and Forward Proof}
+\index{theorems|(}
+Theorems, which represent the axioms, theorems and rules of object-logics,
+have type {\tt thm}\indexbold{*thm}.  This chapter begins by describing
+operations that print theorems and that join them in forward proof.  Most
+theorem operations are intended for advanced applications, such as
+programming new proof procedures.  Many of these operations refer to
+signatures, certified terms and certified types, which have the \ML{} types
+{\tt Sign.sg}, {\tt Sign.cterm} and {\tt Sign.ctyp} and are discussed in
+Chapter~\ref{theories}.  Beginning users should ignore such complexities
+--- and skip all but the first section of this chapter.
+
+The theorem operations do not print error messages.  Instead, they raise
+exception~\ttindex{THM}\@.  Use \ttindex{print_exn} to display
+exceptions nicely:
+\begin{ttbox} 
+allI RS mp  handle e => print_exn e;
+{\out Exception THM raised:}
+{\out RSN: no unifiers -- premise 1}
+{\out (!!x. ?P(x)) ==> ALL x. ?P(x)}
+{\out [| ?P --> ?Q; ?P |] ==> ?Q}
+{\out}
+{\out uncaught exception THM}
+\end{ttbox}
+
+
+\section{Basic operations on theorems}
+\subsection{Pretty-printing a theorem}
+\index{theorems!printing|bold}\index{printing!theorems|bold}
+\subsubsection{Top-level commands}
+\begin{ttbox} 
+prth: thm -> thm
+prths: thm list -> thm list
+prthq: thm Sequence.seq -> thm Sequence.seq
+\end{ttbox}
+These are for interactive use.  They are identity functions that display,
+then return, their argument.  The \ML{} identifier {\tt it} will refer to
+the value just displayed.
+\begin{description}
+\item[\ttindexbold{prth} {\it thm}]  
+prints {\it thm\/} at the terminal.
+
+\item[\ttindexbold{prths} {\it thms}]  
+prints {\it thms}, a list of theorems.
+
+\item[\ttindexbold{prthq} {\it thmq}]  
+prints {\it thmq}, a sequence of theorems.  It is useful for inspecting
+the output of a tactic.
+\end{description}
+
+\subsubsection{Operations for programming}
+\begin{ttbox} 
+print_thm     : thm -> unit
+print_goals   : int -> thm -> unit
+string_of_thm : thm -> string
+\end{ttbox}
+Functions with result type {\tt unit} are convenient for imperative
+programming.
+\begin{description}
+\item[\ttindexbold{print_thm} {\it thm}]  
+prints {\it thm\/} at the terminal.
+
+\item[\ttindexbold{print_goals} {\it limit\/} {\it thm}]  
+prints {\it thm\/} in goal style, with the premises as subgoals.  It prints
+at most {\it limit\/} subgoals.  The subgoal module calls {\tt print_goals}
+to display proof states.
+
+\item[\ttindexbold{string_of_thm} {\it thm}]  
+converts {\it thm\/} to a string.
+\end{description}
+
+
+\subsection{Forwards proof: joining rules by resolution}
+\index{theorems!joining by resolution|bold}
+\index{meta-rules!resolution|bold}
+\begin{ttbox} 
+RSN : thm * (int * thm) -> thm                 \hfill{\bf infix}
+RS  : thm * thm -> thm                         \hfill{\bf infix}
+MRS : thm list * thm -> thm                    \hfill{\bf infix}
+RLN : thm list * (int * thm list) -> thm list  \hfill{\bf infix}
+RL  : thm list * thm list -> thm list          \hfill{\bf infix}
+MRL: thm list list * thm list -> thm list      \hfill{\bf infix}
+\end{ttbox}
+Putting rules together is a simple way of deriving new rules.  These
+functions are especially useful with destruction rules.
+\begin{description}
+\item[\tt$thm@1$ RSN $(i,thm@2)$] \indexbold{*RSN} 
+resolves the conclusion of $thm@1$ with the $i$th premise of~$thm@2$.  It
+raises exception \ttindex{THM} unless there is precisely one resolvent.
+
+\item[\tt$thm@1$ RS $thm@2$] \indexbold{*RS} 
+abbreviates \hbox{\tt$thm@1$ RSN $(1,thm@2)$}.  Thus, it resolves the
+conclusion of $thm@1$ with the first premise of~$thm@2$.
+
+\item[\tt {$[thm@1,\ldots,thm@n]$} MRS $thm$] \indexbold{*MRS} 
+  uses {\tt RLN} to resolve $thm@i$ against premise~$i$ of $thm$, for
+  $i=n$, \ldots,~1.  This applies $thm@n$, \ldots, $thm@1$ to the first $n$
+  premises of $thm$.  Because the theorems are used from right to left, it
+  does not matter if the $thm@i$ create new premises.  {\tt MRS} is useful
+  for expressing proof trees.
+
+\item[\tt$thms1$ RLN $(i,thms2)$] \indexbold{*RLN} 
+for every $thm@1$ in $thms1$ and $thm@2$ in $thms2$, it resolves the
+conclusion of $thm@1$ with the $i$th premise of~$thm@2$, accumulating the
+results.  It is useful for combining lists of theorems.
+
+\item[\tt$thms1$ RL $thms2$] \indexbold{*RL} 
+abbreviates \hbox{\tt$thms1$ RLN $(1,thms2)$}. 
+
+\item[\tt {$[thms@1,\ldots,thms@n]$} MRL $thms$] \indexbold{*MRL} 
+is analogous to {\tt MRS}, but combines theorem lists rather than theorems.
+It too is useful for expressing proof trees.
+\end{description}
+
+
+\subsection{Expanding definitions in theorems}
+\index{theorems!meta-level rewriting in|bold}\index{rewriting!meta-level}
+\begin{ttbox} 
+rewrite_rule       : thm list -> thm -> thm
+rewrite_goals_rule : thm list -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{rewrite_rule} {\it defs} {\it thm}]  
+unfolds the {\it defs} throughout the theorem~{\it thm}.
+
+\item[\ttindexbold{rewrite_goals_rule} {\it defs} {\it thm}]  
+unfolds the {\it defs} in the premises of~{\it thm}, but leaves the
+conclusion unchanged.  This rule underlies \ttindex{rewrite_goals_tac}, but 
+serves little purpose in forward proof.
+\end{description}
+
+
+\subsection{Instantiating a theorem} \index{theorems!instantiating|bold}
+\begin{ttbox} 
+read_instantiate    :            (string*string)list -> thm -> thm
+read_instantiate_sg : Sign.sg -> (string*string)list -> thm -> thm
+cterm_instantiate   :    (Sign.cterm*Sign.cterm)list -> thm -> thm
+\end{ttbox}
+These meta-rules instantiate type and term unknowns in a theorem.  They are
+occasionally useful.  They can prevent difficulties with higher-order
+unification, and define specialized versions of rules.
+\begin{description}
+\item[\ttindexbold{read_instantiate} {\it insts} {\it thm}] 
+processes the instantiations {\it insts} and instantiates the rule~{\it
+thm}.  The processing of instantiations is described
+in~\S\ref{res_inst_tac}, under {\tt res_inst_tac}.  
+
+Use {\tt res_inst_tac}, not {\tt read_instantiate}, to instantiate a rule
+and refine a particular subgoal.  The tactic allows instantiation by the
+subgoal's parameters, and reads the instantiations using the signature
+associated with the proof state.  The remaining two instantiation functions
+are highly specialized; most users should ignore them.
+
+\item[\ttindexbold{read_instantiate_sg} {\it sg} {\it insts} {\it thm}] 
+resembles \hbox{\tt read_instantiate {\it insts} {\it thm}}, but reads the
+instantiates under signature~{\it sg}.  This is necessary when you want to
+instantiate a rule from a general theory, such as first-order logic, using
+the notation of some specialized theory.  Use the function {\tt
+sign_of} to get the signature of a theory.
+
+\item[\ttindexbold{cterm_instantiate} {\it ctpairs} {\it thm}] 
+is similar to {\tt read_instantiate}, but the instantiations are provided
+as pairs of certified terms, not as strings to be read.
+\end{description}
+
+
+\subsection{Miscellaneous forward rules}
+\index{theorems!standardizing|bold}
+\begin{ttbox} 
+standard         : thm -> thm
+zero_var_indexes : thm -> thm
+make_elim        : thm -> thm
+rule_by_tactic   : tactic -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{standard} $thm$]  
+puts $thm$ into the standard form of object-rules.  It discharges all
+meta-hypotheses, replaces free variables by schematic variables, and
+renames schematic variables to have subscript zero.
+
+\item[\ttindexbold{zero_var_indexes} $thm$] 
+makes all schematic variables have subscript zero, renaming them to avoid
+clashes. 
+
+\item[\ttindexbold{make_elim} $thm$] 
+\index{rules!converting destruction to elimination}
+converts $thm$, a destruction rule of the form $\List{P@1;\ldots;P@m}\Imp
+Q$, to the elimination rule $\List{P@1; \ldots; P@m; Q\Imp R}\Imp R$.  This
+is the basis for destruct-resolution: {\tt dresolve_tac}, etc.
+
+\item[\ttindexbold{rule_by_tactic} {\it tac} {\it thm}] 
+  applies {\it tac\/} to the {\it thm}, freezing its variables first, then
+  yields the proof state returned by the tactic.  In typical usage, the
+  {\it thm\/} represents an instance of a rule with several premises, some
+  with contradictory assumptions (because of the instantiation).  The
+  tactic proves those subgoals and does whatever else it can, and returns
+  whatever is left.
+\end{description}
+
+
+\subsection{Taking a theorem apart}
+\index{theorems!taking apart|bold}
+\index{flex-flex constraints}
+\begin{ttbox} 
+concl_of      : thm -> term
+prems_of      : thm -> term list
+nprems_of     : thm -> int
+tpairs_of     : thm -> (term*term)list
+stamps_of_thy : thm -> string ref list
+dest_state    : thm*int -> (term*term)list * term list * term * term
+rep_thm       : thm -> \{prop:term, hyps:term list, 
+                        maxidx:int, sign:Sign.sg\}
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{concl_of} $thm$] 
+returns the conclusion of $thm$ as a term.
+
+\item[\ttindexbold{prems_of} $thm$] 
+returns the premises of $thm$ as a list of terms.
+
+\item[\ttindexbold{nprems_of} $thm$] 
+returns the number of premises in $thm$: {\tt length(prems_of~$thm$)}.
+
+\item[\ttindexbold{tpairs_of} $thm$] 
+returns the flex-flex constraints of $thm$.
+
+\item[\ttindexbold{stamps_of_thm} $thm$] 
+returns the stamps of the signature associated with~$thm$.
+
+\item[\ttindexbold{dest_state} $(thm,i)$] 
+decomposes $thm$ as a tuple containing a list of flex-flex constraints, a
+list of the subgoals~1 to~$i-1$, subgoal~$i$, and the rest of the theorem
+(this will be an implication if there are more than $i$ subgoals).
+
+\item[\ttindexbold{rep_thm} $thm$] 
+decomposes $thm$ as a record containing the statement of~$thm$, its list of
+meta-hypotheses, the maximum subscript of its unknowns, and its signature.
+\end{description}
+
+
+\subsection{Tracing flags for unification}
+\index{tracing!of unification}
+\begin{ttbox} 
+Unify.trace_simp   : bool ref \hfill{\bf initially false}
+Unify.trace_types  : bool ref \hfill{\bf initially false}
+Unify.trace_bound  : int ref \hfill{\bf initially 10}
+Unify.search_bound : int ref \hfill{\bf initially 20}
+\end{ttbox}
+Tracing the search may be useful when higher-order unification behaves
+unexpectedly.  Letting {\tt res_inst_tac} circumvent the problem is easier,
+though.
+\begin{description}
+\item[\ttindexbold{Unify.trace_simp} \tt:= true;] 
+causes tracing of the simplification phase.
+
+\item[\ttindexbold{Unify.trace_types} \tt:= true;] 
+generates warnings of incompleteness, when unification is not considering
+all possible instantiations of type unknowns.
+
+\item[\ttindexbold{Unify.trace_bound} \tt:= $n$] 
+causes unification to print tracing information once it reaches depth~$n$.
+Use $n=0$ for full tracing.  At the default value of~10, tracing
+information is almost never printed.
+
+\item[\ttindexbold{Unify.search_bound} \tt:= $n$] 
+causes unification to limit its search to depth~$n$.  Because of this
+bound, higher-order unification cannot return an infinite sequence, though
+it can return a very long one.  The search rarely approaches the default
+value of~20.  If the search is cut off, unification prints {\tt
+***Unification bound exceeded}.
+\end{description}
+
+
+\section{Primitive meta-level inference rules}
+\index{meta-rules|(}
+These implement the meta-logic in {\sc lcf} style, as functions from theorems
+to theorems.  They are, rarely, useful for deriving results in the pure
+theory.  Mainly, they are included for completeness, and most users should
+not bother with them.  The meta-rules raise exception \ttindex{THM} to signal
+malformed premises, incompatible signatures and similar errors.
+
+The meta-logic uses natural deduction.  Each theorem may depend on
+meta-hypotheses, or assumptions.  Certain rules, such as $({\Imp}I)$,
+discharge assumptions; in most other rules, the conclusion depends on all
+of the assumptions of the premises.  Formally, the system works with
+assertions of the form
+\[ \phi \quad [\phi@1,\ldots,\phi@n], \]
+where $\phi@1$,~\ldots,~$\phi@n$ are the hypotheses.  Do not confuse
+meta-level assumptions with the object-level assumptions in a subgoal,
+which are represented in the meta-logic using~$\Imp$.
+
+Each theorem has a signature.  Certified terms have a signature.  When a
+rule takes several premises and certified terms, it merges the signatures
+to make a signature for the conclusion.  This fails if the signatures are
+incompatible. 
+
+The {\em implication\/} rules are $({\Imp}I)$
+and $({\Imp}E)$:
+\[ \infer[({\Imp}I)]{\phi\Imp \psi}{\infer*{\psi}{[\phi]}}  \qquad
+   \infer[({\Imp}E)]{\psi}{\phi\Imp \psi & \phi}  \]
+
+Equality of truth values means logical equivalence:
+\[ \infer[({\equiv}I)]{\phi\equiv\psi}{\infer*{\psi}{[\phi]} &
+ 			               \infer*{\phi}{[\psi]}}  
+   \qquad
+   \infer[({\equiv}E)]{\psi}{\phi\equiv \psi & \phi}   \]
+
+The {\em equality\/} rules are reflexivity, symmetry, and transitivity:
+\[ {a\equiv a}\,(refl)  \qquad
+   \infer[(sym)]{b\equiv a}{a\equiv b}  \qquad
+   \infer[(trans)]{a\equiv c}{a\equiv b & b\equiv c}   \]
+
+The $\lambda$-conversions are $\alpha$-conversion, $\beta$-conversion, and
+extensionality:\footnote{$\alpha$-conversion holds if $y$ is not free
+in~$a$; $(ext)$ holds if $x$ is not free in the assumptions, $f$, or~$g$.}
+\[ {(\lambda x.a) \equiv (\lambda y.a[y/x])}    \qquad
+   {((\lambda x.a)(b)) \equiv a[b/x]}           \qquad
+   \infer[(ext)]{f\equiv g}{f(x) \equiv g(x)}   \]
+
+The {\it abstraction\/} and {\it combination\/} rules permit the conversion
+of subterms:\footnote{Abstraction holds if $x$ is not free in the
+assumptions.}
+\[  \infer[(abs)]{(\lambda x.a) \equiv (\lambda x.b)}{a\equiv b}   \qquad
+    \infer[(comb)]{f(a)\equiv g(b)}{f\equiv g & a\equiv b}   \]
+
+The {\em universal quantification\/} rules are $(\Forall I)$ and $(\Forall
+E)$:\footnote{$(\Forall I)$ holds if $x$ is not free in the assumptions.}
+\[ \infer[(\Forall I)]{\Forall x.\phi}{\phi}        \qquad
+   \infer[(\Forall I)]{\phi[b/x]}{\Forall x.\phi}   \]
+
+
+\subsection{Propositional rules}
+\index{meta-rules!assumption|bold}
+\subsubsection{Assumption}
+\begin{ttbox} 
+assume: Sign.cterm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{assume} $ct$] 
+makes the theorem \(\phi \quad[\phi]\), where $\phi$ is the value of~$ct$.
+The rule checks that $ct$ has type $prop$ and contains no unknowns, which
+are not allowed in hypotheses.
+\end{description}
+
+\subsubsection{Implication}
+\index{meta-rules!implication|bold}
+\begin{ttbox} 
+implies_intr      : Sign.cterm -> thm -> thm
+implies_intr_list : Sign.cterm list -> thm -> thm
+implies_intr_hyps : thm -> thm
+implies_elim      : thm -> thm -> thm
+implies_elim_list : thm -> thm list -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{implies_intr} $ct$ $thm$] 
+is $({\Imp}I)$, where $ct$ is the assumption to discharge, say~$\phi$.  It
+maps the premise $\psi\quad[\phi]$ to the conclusion $\phi\Imp\psi$.  The
+rule checks that $ct$ has type $prop$.
+
+\item[\ttindexbold{implies_intr_list} $cts$ $thm$] 
+applies $({\Imp}I)$ repeatedly, on every element of the list~$cts$.
+
+\item[\ttindexbold{implies_intr_hyps} $thm$] 
+applies $({\Imp}I)$ to discharge all the hypotheses of~$thm$.  It maps the
+premise $\phi \quad [\phi@1,\ldots,\phi@n]$ to the conclusion
+$\List{\phi@1,\ldots,\phi@n}\Imp\phi$.
+
+\item[\ttindexbold{implies_elim} $thm@1$ $thm@2$] 
+applies $({\Imp}E)$ to $thm@1$ and~$thm@2$.  It maps the premises $\phi\Imp
+\psi$ and $\phi$ to the conclusion~$\psi$.
+
+\item[\ttindexbold{implies_elim_list} $thm$ $thms$] 
+applies $({\Imp}E)$ repeatedly to $thm$, using each element of~$thms$ in
+turn.  It maps the premises $[\phi@1,\ldots,\phi@n]\Imp\psi$ and
+$\phi@1$,\ldots,$\phi@n$ to the conclusion~$\psi$.
+\end{description}
+
+\subsubsection{Logical equivalence (equality)}
+\index{meta-rules!logical equivalence|bold}
+\begin{ttbox} 
+equal_intr : thm -> thm -> thm equal_elim : thm -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{equal_intr} $thm@1$ $thm@2$] 
+applies $({\equiv}I)$ to $thm@1$ and~$thm@2$.  It maps the premises
+$\psi\quad[\phi]$ and $\phi\quad[\psi]$ to the conclusion~$\phi\equiv\psi$.
+
+\item[\ttindexbold{equal_elim} $thm@1$ $thm@2$] 
+applies $({\equiv}E)$ to $thm@1$ and~$thm@2$.  It maps the premises
+$\phi\equiv\psi$ and $\phi$ to the conclusion~$\psi$.
+\end{description}
+
+
+\subsection{Equality rules}
+\index{meta-rules!equality|bold}
+\begin{ttbox} 
+reflexive  : Sign.cterm -> thm
+symmetric  : thm -> thm
+transitive : thm -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{reflexive} $ct$] 
+makes the theorem \(ct=ct\). 
+
+\item[\ttindexbold{symmetric} $thm$] 
+maps the premise $a\equiv b$ to the conclusion $b\equiv a$.
+
+\item[\ttindexbold{transitive} $thm@1$ $thm@2$] 
+maps the premises $a\equiv b$ and $b\equiv c$ to the conclusion~${a\equiv c}$.
+\end{description}
+
+
+\subsection{The $\lambda$-conversion rules}
+\index{meta-rules!$\lambda$-conversion|bold}
+\begin{ttbox} 
+beta_conversion : Sign.cterm -> thm
+extensional     : thm -> thm
+abstract_rule   : string -> Sign.cterm -> thm -> thm
+combination     : thm -> thm -> thm
+\end{ttbox} 
+There is no rule for $\alpha$-conversion because Isabelle's internal
+representation ignores bound variable names, except when communicating with
+the user.
+\begin{description}
+\item[\ttindexbold{beta_conversion} $ct$] 
+makes the theorem $((\lambda x.a)(b)) \equiv a[b/x]$, where $ct$ is the
+term $(\lambda x.a)(b)$.
+
+\item[\ttindexbold{extensional} $thm$] 
+maps the premise $f(x) \equiv g(x)$ to the conclusion $f\equiv g$.
+Parameter~$x$ is taken from the premise.  It may be an unknown or a free
+variable (provided it does not occur in the hypotheses); it must not occur
+in $f$ or~$g$.
+
+\item[\ttindexbold{abstract_rule} $v$ $x$ $thm$] 
+maps the premise $a\equiv b$ to the conclusion $(\lambda x.a) \equiv
+(\lambda x.b)$, abstracting over all occurrences (if any!) of~$x$.
+Parameter~$x$ is supplied as a cterm.  It may be an unknown or a free
+variable (provided it does not occur in the hypotheses).  In the
+conclusion, the bound variable is named~$v$.
+
+\item[\ttindexbold{combination} $thm@1$ $thm@2$] 
+maps the premises $f\equiv g$ and $a\equiv b$ to the conclusion~$f(a)\equiv
+g(b)$.
+\end{description}
+
+
+\subsection{Universal quantifier rules}
+\index{meta-rules!quantifier|bold}
+\subsubsection{Forall introduction}
+\begin{ttbox} 
+forall_intr       : Sign.cterm      -> thm -> thm
+forall_intr_list  : Sign.cterm list -> thm -> thm
+forall_intr_frees :                    thm -> thm
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{forall_intr} $x$ $thm$] 
+applies $({\Forall}I)$, abstracting over all occurrences (if any!) of~$x$.
+The rule maps the premise $\phi$ to the conclusion $\Forall x.\phi$.
+Parameter~$x$ is supplied as a cterm.  It may be an unknown or a free
+variable (provided it does not occur in the hypotheses).
+
+\item[\ttindexbold{forall_intr_list} $xs$ $thm$] 
+applies $({\Forall}I)$ repeatedly, on every element of the list~$xs$.
+
+\item[\ttindexbold{forall_intr_frees} $thm$] 
+applies $({\Forall}I)$ repeatedly, generalizing over all the free variables
+of the premise.
+\end{description}
+
+
+\subsubsection{Forall elimination}
+\begin{ttbox} 
+forall_elim       : Sign.cterm      -> thm -> thm
+forall_elim_list  : Sign.cterm list -> thm -> thm
+forall_elim_var   :             int -> thm -> thm
+forall_elim_vars  :             int -> thm -> thm
+\end{ttbox}
+
+\begin{description}
+\item[\ttindexbold{forall_elim} $ct$ $thm$] 
+applies $({\Forall}E)$, mapping the premise $\Forall x.\phi$ to the conclusion
+$\phi[ct/x]$.  The rule checks that $ct$ and $x$ have the same type.
+
+\item[\ttindexbold{forall_elim_list} $cts$ $thm$] 
+applies $({\Forall}E)$ repeatedly, on every element of the list~$cts$.
+
+\item[\ttindexbold{forall_elim_var} $k$ $thm$] 
+applies $({\Forall}E)$, mapping the premise $\Forall x.\phi$ to the conclusion
+$\phi[\Var{x@k}/x]$.  Thus, it replaces the outermost $\Forall$-bound
+variable by an unknown having subscript~$k$.
+
+\item[\ttindexbold{forall_elim_vars} $ks$ $thm$] 
+applies {\tt forall_elim_var} repeatedly, for every element of the list~$ks$.
+\end{description}
+
+\subsubsection{Instantiation of unknowns}
+\index{meta-rules!instantiation|bold}
+\begin{ttbox} 
+instantiate: (indexname*Sign.ctyp)list * (Sign.cterm*Sign.cterm)list
+                   -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{instantiate} $tyinsts$ $insts$ $thm$] 
+performs simultaneous substitution of types for type unknowns (the
+$tyinsts$) and terms for term unknowns (the $insts$).  Instantiations are
+given as $(v,t)$ pairs, where $v$ is an unknown and $t$ is a term (of the
+same type as $v$) or a type (of the same sort as~$v$).  All the unknowns
+must be distinct.  The rule normalizes its conclusion.
+\end{description}
+
+
+\subsection{Freezing/thawing type variables}
+\index{meta-rules!for type variables|bold}
+\begin{ttbox} 
+freezeT: thm -> thm
+varifyT: thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{freezeT} $thm$] 
+converts all the type unknowns in $thm$ to free type variables.
+
+\item[\ttindexbold{varifyT} $thm$] 
+converts all the free type variables in $thm$ to type unknowns.
+\end{description}
+
+
+\section{Derived rules for goal-directed proof}
+Most of these rules have the sole purpose of implementing particular
+tactics.  There are few occasions for applying them directly to a theorem.
+
+\subsection{Proof by assumption}
+\index{meta-rules!assumption|bold}
+\begin{ttbox} 
+assumption    : int -> thm -> thm Sequence.seq
+eq_assumption : int -> thm -> thm
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{assumption} {\it i} $thm$] 
+attempts to solve premise~$i$ of~$thm$ by assumption.
+
+\item[\ttindexbold{eq_assumption}] 
+is like {\tt assumption} but does not use unification.
+\end{description}
+
+
+\subsection{Resolution}
+\index{meta-rules!resolution|bold}
+\begin{ttbox} 
+biresolution : bool -> (bool*thm)list -> int -> thm
+               -> thm Sequence.seq
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{biresolution} $match$ $rules$ $i$ $state$] 
+performs bi-resolution on subgoal~$i$ of~$state$, using the list of $\it
+(flag,rule)$ pairs.  For each pair, it applies resolution if the flag
+is~{\tt false} and elim-resolution if the flag is~{\tt true}.  If $match$
+is~{\tt true}, the $state$ is not instantiated.
+\end{description}
+
+
+\subsection{Composition: resolution without lifting}
+\index{meta-rules!for composition|bold}
+\begin{ttbox}
+compose   : thm * int * thm -> thm list
+COMP      : thm * thm -> thm
+bicompose : bool -> bool * thm * int -> int -> thm
+            -> thm Sequence.seq
+\end{ttbox}
+In forward proof, a typical use of composition is to regard an assertion of
+the form $\phi\Imp\psi$ as atomic.  Schematic variables are not renamed, so
+beware of clashes!
+\begin{description}
+\item[\ttindexbold{compose} ($thm@1$, $i$, $thm@2$)] 
+uses $thm@1$, regarded as an atomic formula, to solve premise~$i$
+of~$thm@2$.  Let $thm@1$ and $thm@2$ be $\psi$ and $\List{\phi@1; \ldots;
+\phi@n} \Imp \phi$.  For each $s$ that unifies~$\psi$ and $\phi@i$, the
+result list contains the theorem
+\[ (\List{\phi@1; \ldots; \phi@{i-1}; \phi@{i+1}; \ldots; \phi@n} \Imp \phi)s.
+\]
+
+\item[\tt $thm@1$ COMP $thm@2$] 
+calls \hbox{\tt compose ($thm@1$, 1, $thm@2$)} and returns the result, if
+unique; otherwise, it raises exception~\ttindex{THM}\@.  It is
+analogous to {\tt RS}\@.  
+
+For example, suppose that $thm@1$ is $a=b\Imp b=a$, a symmetry rule, and
+that $thm@2$ is $\List{P\Imp Q; \neg Q} \Imp\neg A)$, which is the
+principle of contrapositives.  Then the result would be the
+derived rule $\neg(b=a)\Imp\neg(a=b)$.
+
+\item[\ttindexbold{bicompose} $match$ ($flag$, $rule$, $m$) $i$ $state$]
+refines subgoal~$i$ of $state$ using $rule$, without lifting.  The $rule$
+is taken to have the form $\List{\psi@1; \ldots; \psi@m} \Imp \psi$, where
+$\psi$ need {\bf not} be atomic; thus $m$ determines the number of new
+subgoals.  If $flag$ is {\tt true} then it performs elim-resolution --- it
+solves the first premise of~$rule$ by assumption and deletes that
+assumption.  If $match$ is~{\tt true}, the $state$ is not instantiated.
+\end{description}
+
+
+\subsection{Other meta-rules}
+\begin{ttbox} 
+trivial            : Sign.cterm -> thm
+lift_rule          : (thm * int) -> thm -> thm
+rename_params_rule : string list * int -> thm -> thm
+rewrite_cterm      : thm list -> Sign.cterm -> thm
+flexflex_rule      : thm -> thm Sequence.seq
+\end{ttbox}
+\begin{description}
+\item[\ttindexbold{trivial} $ct$] 
+makes the theorem \(\phi\Imp\phi\), where $\phi$ is the value of~$ct$.
+This is the initial state for a goal-directed proof of~$\phi$.  The rule
+checks that $ct$ has type~$prop$.
+
+\item[\ttindexbold{lift_rule} ($state$, $i$) $rule$] \index{lifting}
+prepares $rule$ for resolution by lifting it over the parameters and
+assumptions of subgoal~$i$ of~$state$.
+
+\item[\ttindexbold{rename_params_rule} ({\it names}, {\it i}) $thm$] 
+uses the $names$ to rename the parameters of premise~$i$ of $thm$.  The
+names must be distinct.  If there are fewer names than parameters, then the
+rule renames the innermost parameters and may modify the remaining ones to
+ensure that all the parameters are distinct.
+\index{parameters!renaming}
+
+\item[\ttindexbold{rewrite_cterm} $defs$ $ct$]
+transforms $ct$ to $ct'$ by repeatedly applying $defs$ as rewrite rules; it
+returns the conclusion~$ct\equiv ct'$.  This underlies the meta-rewriting
+tactics and rules.
+\index{terms!meta-level rewriting in|bold}\index{rewriting!meta-level}
+
+\item[\ttindexbold{flexflex_rule} $thm$]  \index{flex-flex constraints}
+\index{meta-rules!for flex-flex constraints|bold}
+removes all flex-flex pairs from $thm$ using the trivial unifier.
+\end{description}
+\index{theorems|)}
+\index{meta-rules|)}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/Ref/undocumented.tex	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,319 @@
+%%%%Currently UNDOCUMENTED low-level functions!  from previous manual
+
+%%%%Low level information about terms and module Logic.
+%%%%Mainly used for implementation of Pure.
+
+%move to ML sources?
+
+\subsection{Basic declarations}
+The implication symbol is {\tt implies}.
+
+The term \verb|all T| is the universal quantifier for type {\tt T}\@.
+
+The term \verb|equals T| is the equality predicate for type {\tt T}\@.
+
+
+There are a number of basic functions on terms and types.
+
+\index{--->}
+\beginprog
+op ---> : typ list * typ -> typ
+\endprog
+Given types \([ \tau_1, \ldots, \tau_n]\) and \(\tau\), it
+forms the type \(\tau_1\to \cdots \to (\tau_n\to\tau)\).
+
+\index{loose_bnos}
+\beginprog
+loose_bnos: term -> int list
+\endprog
+Calling \verb|loose_bnos t| returns the list of all 'loose' bound variable
+references.  In particular, {\tt Bound~0} is loose unless it is enclosed in
+an abstraction.  Similarly {\tt Bound~1} is loose unless it is enclosed in
+at least two abstractions; if enclosed in just one, the list will contain
+the number 0.  A well-formed term does not contain any loose variables.
+
+Calling {\prog{}type_of \${t}}\index{type_of} computes the type of the
+term~$t$.  Raises exception {\tt TYPE} unless applications are well-typed.
+
+\index{aconv}
+\beginprog
+op aconv: term*term -> bool
+\endprog
+Calling $t\Term{ aconv }u$ tests whether terms~$t$ and~$u$ are
+\(\alpha\)-convertible: identical up to renaming of bound variables.
+\begin{itemize}
+  \item
+    Two constants, {\tt Free}s, or {\tt Var}s are \(\alpha\)-convertible
+    just when their names and types are equal.
+    (Variables having the same name but different types are thus distinct.
+    This confusing situation should be avoided!)
+  \item
+    Two bound variables are \(\alpha\)-convertible
+    just when they have the same number.
+  \item
+    Two abstractions are \(\alpha\)-convertible
+    just when their bodies are, and their bound variables have the same type.
+  \item
+    Two applications are \(\alpha\)-convertible
+    just when the corresponding subterms are.
+\end{itemize}
+
+
+\index{incr_boundvars}
+\beginprog
+incr_boundvars: int -> term -> term
+\endprog
+This increments a term's `loose' bound variables by a given offset,
+required when moving a subterm into a context
+where it is enclosed by a different number of lambdas.
+
+\index{abstract_over}
+\beginprog
+abstract_over: term*term -> term
+\endprog
+For abstracting a term over a subterm \(v\):
+replaces every occurrence of \(v\) by a {\tt Bound} variable
+with the correct index.
+
+Calling \verb|subst_bounds|$([u_{n-1},\ldots,u_0],\,t)$\index{subst_bounds}
+substitutes the $u_i$ for loose bound variables in $t$.  This achieves
+\(\beta\)-reduction of \(u_{n-1} \cdots u_0\) into $t$, replacing {\tt
+Bound~i} with $u_i$.  For \((\lambda x y.t)(u,v)\), the bound variable
+indices in $t$ are $x:1$ and $y:0$.  The appropriate call is
+\verb|subst_bounds([v,u],t)|.  Loose bound variables $\geq n$ are reduced
+by $n$ to compensate for the disappearance of $n$ lambdas.
+
+\index{maxidx_of_term}
+\beginprog
+maxidx_of_term: term -> int
+\endprog
+Computes the maximum index of all the {\tt Var}s in a term.
+If there are no {\tt Var}s, the result is \(-1\).
+
+\index{term_match}
+\beginprog
+term_match: (term*term)list * term*term -> (term*term)list
+\endprog
+Calling \verb|term_match(vts,t,u)| instantiates {\tt Var}s in {\tt t} to
+match it with {\tt u}.  The resulting list of variable/term pairs extends
+{\tt vts}, which is typically empty.  First-order pattern matching is used
+to implement meta-level rewriting.
+
+
+\subsection{The representation of object-rules}
+The module {\tt Logic} contains operations concerned with inference ---
+especially, for constructing and destructing terms that represent
+object-rules.
+
+\index{occs}
+\beginprog
+op occs: term*term -> bool
+\endprog
+Does one term occur in the other?
+(This is a reflexive relation.)
+
+\index{add_term_vars}
+\beginprog
+add_term_vars: term*term list -> term list
+\endprog
+Accumulates the {\tt Var}s in the term, suppressing duplicates.
+The second argument should be the list of {\tt Var}s found so far.
+
+\index{add_term_frees}
+\beginprog
+add_term_frees: term*term list -> term list
+\endprog
+Accumulates the {\tt Free}s in the term, suppressing duplicates.
+The second argument should be the list of {\tt Free}s found so far.
+
+\index{mk_equals}
+\beginprog
+mk_equals: term*term -> term
+\endprog
+Given $t$ and $u$ makes the term $t\equiv u$.
+
+\index{dest_equals}
+\beginprog
+dest_equals: term -> term*term
+\endprog
+Given $t\equiv u$ returns the pair $(t,u)$.
+
+\index{list_implies:}
+\beginprog
+list_implies: term list * term -> term
+\endprog
+Given the pair $([\phi_1,\ldots, \phi_m], \phi)$
+makes the term \([\phi_1;\ldots; \phi_m] \Imp \phi\).
+
+\index{strip_imp_prems}
+\beginprog
+strip_imp_prems: term -> term list
+\endprog
+Given \([\phi_1;\ldots; \phi_m] \Imp \phi\)
+returns the list \([\phi_1,\ldots, \phi_m]\). 
+
+\index{strip_imp_concl}
+\beginprog
+strip_imp_concl: term -> term
+\endprog
+Given \([\phi_1;\ldots; \phi_m] \Imp \phi\)
+returns the term \(\phi\). 
+
+\index{list_equals}
+\beginprog
+list_equals: (term*term)list * term -> term
+\endprog
+For adding flex-flex constraints to an object-rule. 
+Given $([(t_1,u_1),\ldots, (t_k,u_k)], \phi)$,
+makes the term \([t_1\equiv u_1;\ldots; t_k\equiv u_k]\Imp \phi\).
+
+\index{strip_equals}
+\beginprog
+strip_equals: term -> (term*term) list * term
+\endprog
+Given \([t_1\equiv u_1;\ldots; t_k\equiv u_k]\Imp \phi\),
+returns $([(t_1,u_1),\ldots, (t_k,u_k)], \phi)$.
+
+\index{rule_of}
+\beginprog
+rule_of: (term*term)list * term list * term -> term
+\endprog
+Makes an object-rule: given the triple
+\[ ([(t_1,u_1),\ldots, (t_k,u_k)], [\phi_1,\ldots, \phi_m], \phi) \]
+returns the term
+\([t_1\equiv u_1;\ldots; t_k\equiv u_k; \phi_1;\ldots; \phi_m]\Imp \phi\)
+
+\index{strip_horn}
+\beginprog
+strip_horn: term -> (term*term)list * term list * term
+\endprog
+Breaks an object-rule into its parts: given
+\[ [t_1\equiv u_1;\ldots; t_k\equiv u_k; \phi_1;\ldots; \phi_m] \Imp \phi \]
+returns the triple
+\(([(t_k,u_k),\ldots, (t_1,u_1)], [\phi_1,\ldots, \phi_m], \phi).\)
+
+\index{strip_assums}
+\beginprog
+strip_assums: term -> (term*int) list * (string*typ) list * term
+\endprog
+Strips premises of a rule allowing a more general form,
+where $\Forall$ and $\Imp$ may be intermixed.
+This is typical of assumptions of a subgoal in natural deduction.
+Returns additional information about the number, names,
+and types of quantified variables.
+
+
+\index{strip_prems}
+\beginprog
+strip_prems: int * term list * term -> term list * term
+\endprog
+For finding premise (or subgoal) $i$: given the triple
+\( (i, [], \phi_1;\ldots \phi_i\Imp \phi) \)
+it returns another triple,
+\((\phi_i, [\phi_{i-1},\ldots, \phi_1], \phi)\),
+where $\phi$ need not be atomic. Raises an exception if $i$ is out of
+range.
+
+
+\subsection{Environments}
+The module {\tt Envir} (which is normally closed)
+declares a type of environments.
+An environment holds variable assignments
+and the next index to use when generating a variable.
+\par\indent\vbox{\small \begin{verbatim}
+    datatype env = Envir of {asol: term xolist, maxidx: int}
+\end{verbatim}}
+The operations of lookup, update, and generation of variables
+are used during unification.
+
+\beginprog
+empty: int->env
+\endprog
+Creates the environment with no assignments
+and the given index.
+
+\beginprog
+lookup: env * indexname -> term option
+\endprog
+Looks up a variable, specified by its indexname,
+and returns {\tt None} or {\tt Some} as appropriate.
+
+\beginprog
+update: (indexname * term) * env -> env
+\endprog
+Given a variable, term, and environment,
+produces {\em a new environment\/} where the variable has been updated.
+This has no side effect on the given environment.
+
+\beginprog
+genvar: env * typ -> env * term
+\endprog
+Generates a variable of the given type and returns it,
+paired with a new environment (with incremented {\tt maxidx} field).
+
+\beginprog
+alist_of: env -> (indexname * term) list
+\endprog
+Converts an environment into an association list
+containing the assignments.
+
+\beginprog
+norm_term: env -> term -> term
+\endprog
+
+Copies a term, 
+following assignments in the environment,
+and performing all possible \(\beta\)-reductions.
+
+\beginprog
+rewrite: (env * (term*term)list) -> term -> term
+\endprog
+Rewrites a term using the given term pairs as rewrite rules.  Assignments
+are ignored; the environment is used only with {\tt genvar}, to generate
+unique {\tt Var}s as placeholders for bound variables.
+
+
+\subsection{The unification functions}
+
+
+\beginprog
+unifiers: env * ((term*term)list) -> (env * (term*term)list) seq
+\endprog
+This is the main unification function.
+Given an environment and a list of disagreement pairs,
+it returns a sequence of outcomes.
+Each outcome consists of an updated environment and 
+a list of flex-flex pairs (these are discussed below).
+
+\beginprog
+smash_unifiers: env * (term*term)list -> env seq
+\endprog
+This unification function maps an environment and a list of disagreement
+pairs to a sequence of updated environments. The function obliterates
+flex-flex pairs by choosing the obvious unifier. It may be used to tidy up
+any flex-flex pairs remaining at the end of a proof.
+
+\subsubsection{Flexible-flexible disagreement pairs}
+
+
+\subsubsection{Multiple unifiers}
+The unification procedure performs Huet's {\sc match} operation
+\cite{huet75} in big steps.
+It solves \(\Var{f}(t_1,\ldots,t_p) \equiv u\) for \(\Var{f}\) by finding
+all ways of copying \(u\), first trying projection on the arguments
+\(t_i\).  It never copies below any variable in \(u\); instead it returns a
+new variable, resulting in a flex-flex disagreement pair.  
+
+
+\beginprog
+type_assign: cterm -> cterm
+\endprog
+Produces a cterm by updating the signature of its argument
+to include all variable/type assignments.
+Type inference under the resulting signature will assume the
+same type assignments as in the argument.
+This is used in the goal package to give persistence to type assignments
+within each proof. 
+(Contrast with {\sc lcf}'s sticky types \cite[page 148]{paulson-book}.)
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/ind-defs.bbl	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,105 @@
+\begin{thebibliography}{10}
+
+\bibitem{abramsky90}
+Samson Abramsky.
+\newblock The lazy lambda calculus.
+\newblock In David~A. Turner, editor, {\em Resesarch Topics in Functional
+  Programming}, pages 65--116. Addison-Wesley, 1977.
+
+\bibitem{aczel77}
+Peter Aczel.
+\newblock An introduction to inductive definitions.
+\newblock In J.~Barwise, editor, {\em Handbook of Mathematical Logic}, pages
+  739--782. North-Holland, 1977.
+
+\bibitem{aczel88}
+Peter Aczel.
+\newblock {\em Non-Well-Founded Sets}.
+\newblock CSLI, 1988.
+
+\bibitem{bm79}
+Robert~S. Boyer and J~Strother Moore.
+\newblock {\em A Computational Logic}.
+\newblock Academic Press, 1979.
+
+\bibitem{camilleri92}
+J.~Camilleri and T.~F. Melham.
+\newblock Reasoning with inductively defined relations in the {HOL} theorem
+  prover.
+\newblock Technical Report 265, University of Cambridge Computer Laboratory,
+  August 1992.
+
+\bibitem{davey&priestley}
+B.~A. Davey and H.~A. Priestley.
+\newblock {\em Introduction to Lattices and Order}.
+\newblock Cambridge University Press, 1990.
+
+\bibitem{hennessy90}
+Matthew Hennessy.
+\newblock {\em The Semantics of Programming Languages: An Elementary
+  Introduction Using Structural Operational Semantics}.
+\newblock Wiley, 1990.
+
+\bibitem{melham89}
+Thomas~F. Melham.
+\newblock Automating recursive type definitions in higher order logic.
+\newblock In Graham Birtwistle and P.~A. Subrahmanyam, editors, {\em Current
+  Trends in Hardware Verification and Automated Theorem Proving}, pages
+  341--386. Springer, 1989.
+
+\bibitem{milner89}
+Robin Milner.
+\newblock {\em Communication and Concurrency}.
+\newblock Prentice-Hall, 1989.
+
+\bibitem{paulin92}
+Christine Paulin-Mohring.
+\newblock Inductive definitions in the system {Coq}: Rules and properties.
+\newblock Research Report 92-49, LIP, Ecole Normale Sup\'erieure de Lyon,
+  December 1992.
+
+\bibitem{paulson-set-I}
+Lawrence~C. Paulson.
+\newblock Set theory for verification: {I}. {From} foundations to functions.
+\newblock {\em Journal of Automated Reasoning}.
+\newblock In press; draft available as Report 271, University of Cambridge
+  Computer Laboratory.
+
+\bibitem{paulson91}
+Lawrence~C. Paulson.
+\newblock {\em {ML} for the Working Programmer}.
+\newblock Cambridge University Press, 1991.
+
+\bibitem{paulson-coind}
+Lawrence~C. Paulson.
+\newblock Co-induction and co-recursion in higher-order logic.
+\newblock Technical Report 304, University of Cambridge Computer Laboratory,
+  July 1993.
+
+\bibitem{isabelle-intro}
+Lawrence~C. Paulson.
+\newblock Introduction to {Isabelle}.
+\newblock Technical Report 280, University of Cambridge Computer Laboratory,
+  1993.
+
+\bibitem{paulson-set-II}
+Lawrence~C. Paulson.
+\newblock Set theory for verification: {II}. {Induction} and recursion.
+\newblock Technical Report 312, University of Cambridge Computer Laboratory,
+  1993.
+
+\bibitem{pitts94}
+Andrew~M. Pitts.
+\newblock A co-induction principle for recursively defined domains.
+\newblock {\em Theoretical Computer Science (Fundamental Studies)}.
+\newblock In press; available as Report 252, University of Cambridge Computer
+  Laboratory.
+
+\bibitem{szasz93}
+Nora Szasz.
+\newblock A machine checked proof that {Ackermann's} function is not primitive
+  recursive.
+\newblock In {G\'erard} Huet and Gordon Plotkin, editors, {\em Logical
+  Environments}, pages 317--338. Cambridge University Press, 1993.
+
+\end{thebibliography}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/doc-src/ind-defs.toc	Wed Nov 10 05:00:57 1993 +0100
@@ -0,0 +1,33 @@
+\contentsline {section}{\numberline {1}Introduction}{1}
+\contentsline {section}{\numberline {2}Fixedpoint operators}{2}
+\contentsline {section}{\numberline {3}Elements of an inductive or co-inductive definition}{2}
+\contentsline {subsection}{\numberline {3.1}The form of the introduction rules}{2}
+\contentsline {subsection}{\numberline {3.2}The fixedpoint definitions}{3}
+\contentsline {subsection}{\numberline {3.3}Mutual recursion}{4}
+\contentsline {subsection}{\numberline {3.4}Proving the introduction rules}{4}
+\contentsline {subsection}{\numberline {3.5}The elimination rule}{4}
+\contentsline {section}{\numberline {4}Induction and co-induction rules}{5}
+\contentsline {subsection}{\numberline {4.1}The basic induction rule}{5}
+\contentsline {subsection}{\numberline {4.2}Mutual induction}{5}
+\contentsline {subsection}{\numberline {4.3}Co-induction}{6}
+\contentsline {section}{\numberline {5}Examples of inductive and co-inductive definitions}{6}
+\contentsline {subsection}{\numberline {5.1}The finite set operator}{7}
+\contentsline {subsection}{\numberline {5.2}Lists of $n$ elements}{7}
+\contentsline {subsection}{\numberline {5.3}A demonstration of {\ptt mk\unhbox \voidb@x \kern .06em \vbox {\hrule width.3em}cases}}{9}
+\contentsline {subsection}{\numberline {5.4}A co-inductive definition: bisimulations on lazy lists}{9}
+\contentsline {subsection}{\numberline {5.5}The accessible part of a relation}{10}
+\contentsline {subsection}{\numberline {5.6}The primitive recursive functions}{11}
+\contentsline {section}{\numberline {6}Datatypes and co-datatypes}{13}
+\contentsline {subsection}{\numberline {6.1}Constructors and their domain}{13}
+\contentsline {subsection}{\numberline {6.2}The case analysis operator}{14}
+\contentsline {subsection}{\numberline {6.3}Example: lists and lazy lists}{15}
+\contentsline {subsection}{\numberline {6.4}Example: mutual recursion}{16}
+\contentsline {subsection}{\numberline {6.5}A four-constructor datatype}{18}
+\contentsline {subsection}{\numberline {6.6}Proving freeness theorems}{19}
+\contentsline {section}{\numberline {7}Conclusions and future work}{20}
+\contentsline {section}{\numberline {A}Inductive and co-inductive definitions: users guide}{22}
+\contentsline {subsection}{\numberline {A.1}The result structure}{22}
+\contentsline {subsection}{\numberline {A.2}The argument structure}{23}
+\contentsline {section}{\numberline {B}Datatype and co-datatype definitions: users guide}{24}
+\contentsline {subsection}{\numberline {B.1}The result structure}{24}
+\contentsline {subsection}{\numberline {B.2}The argument structure}{24}