--- a/Nat.thy Fri Nov 25 13:35:32 1994 +0100
+++ b/Nat.thy Fri Nov 25 14:21:14 1994 +0100
@@ -10,53 +10,61 @@
Nat = WF +
+(** type ind **)
+
types
ind
- nat
arities
- ind, nat :: term
-
-instance
- nat :: ord
+ ind :: term
consts
Zero_Rep :: "ind"
Suc_Rep :: "ind => ind"
- Nat :: "ind set"
- Rep_Nat :: "nat => ind"
- Abs_Nat :: "ind => nat"
+
+rules
+ (*the axiom of infinity in 2 parts*)
+ inj_Suc_Rep "inj(Suc_Rep)"
+ Suc_Rep_not_Zero_Rep "Suc_Rep(x) ~= Zero_Rep"
+
+
+
+(** type nat **)
+
+(* type definition *)
+
+subtype (Nat)
+ nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))" (lfp_def)
+
+instance
+ nat :: ord
+
+
+(* abstract constants and syntax *)
+
+consts
+ "0" :: "nat" ("0")
Suc :: "nat => nat"
- nat_case :: "['a, nat=>'a, nat] =>'a"
+ nat_case :: "['a, nat => 'a, nat] => 'a"
pred_nat :: "(nat * nat) set"
nat_rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a"
- "0" :: "nat" ("0")
translations
"case p of 0 => a | Suc(y) => b" == "nat_case(a, %y.b, p)"
-rules
- (*the axiom of infinity in 2 parts*)
- inj_Suc_Rep "inj(Suc_Rep)"
- Suc_Rep_not_Zero_Rep "~(Suc_Rep(x) = Zero_Rep)"
- Nat_def "Nat == lfp(%X. {Zero_Rep} Un (Suc_Rep``X))"
- (*faking a type definition...*)
- Rep_Nat "Rep_Nat(n): Nat"
- Rep_Nat_inverse "Abs_Nat(Rep_Nat(n)) = n"
- Abs_Nat_inverse "i: Nat ==> Rep_Nat(Abs_Nat(i)) = i"
- (*defining the abstract constants*)
- Zero_def "0 == Abs_Nat(Zero_Rep)"
- Suc_def "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
- (*nat operations and recursion*)
- nat_case_def "nat_case(a,f,n) == @z. (n=0 --> z=a) \
-\ & (!x. n=Suc(x) --> z=f(x))"
+defs
+ Zero_def "0 == Abs_Nat(Zero_Rep)"
+ Suc_def "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
+
+ (*nat operations and recursion*)
+ nat_case_def "nat_case(a, f, n) == @z. (n=0 --> z=a) \
+\ & (!x. n=Suc(x) --> z=f(x))"
pred_nat_def "pred_nat == {p. ? n. p = <n, Suc(n)>}"
less_def "m<n == <m,n>:trancl(pred_nat)"
le_def "m<=(n::nat) == ~(n<m)"
- nat_rec_def "nat_rec(n,c,d) == wfrec(pred_nat, n, \
-\ nat_case(%g.c, %m g. d(m,g(m))))"
+ nat_rec_def "nat_rec(n, c, d) == wfrec(pred_nat, n, \
+\ nat_case(%g.c, %m g. d(m, g(m))))"
end
-
--- a/Prod.thy Fri Nov 25 13:35:32 1994 +0100
+++ b/Prod.thy Fri Nov 25 14:21:14 1994 +0100
@@ -1,69 +1,66 @@
-(* Title: HOL/Prod.thy
- ID: $Id$
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+(* Title: HOL/Prod.thy
+ ID: Prod.thy,v 1.5 1994/08/19 09:04:27 lcp Exp
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1992 University of Cambridge
-Ordered Pairs and the Cartesian product type
-The unit type
-
-The type definition admits the following unused axiom:
- Abs_Unit_inverse "f: Unit ==> Rep_Unit(Abs_Unit(f)) = f"
+Ordered Pairs and the Cartesian product type.
+The unit type.
*)
-Prod = Fun +
+Prod = Fun +
+
+(** Products **)
+
+(* type definition *)
+
+consts
+ Pair_Rep :: "['a, 'b] => ['a, 'b] => bool"
-types
- ('a,'b) "*" (infixr 20)
- unit
+defs
+ Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)"
-arities
- "*" :: (term,term)term
- unit :: term
+subtype (Prod)
+ ('a, 'b) "*" (infixr 20)
+ = "{f. ? a b. f = Pair_Rep(a::'a, b::'b)}"
+
+
+(* abstract constants and syntax *)
consts
- Pair_Rep :: "['a,'b] => ['a,'b] => bool"
- Prod :: "('a => 'b => bool)set"
- Rep_Prod :: "'a * 'b => ('a => 'b => bool)"
- Abs_Prod :: "('a => 'b => bool) => 'a * 'b"
- fst :: "'a * 'b => 'a"
- snd :: "'a * 'b => 'b"
- split :: "[['a,'b]=>'c, 'a * 'b] => 'c"
- prod_fun :: "['a=>'b, 'c=>'d, 'a*'c] => 'b*'d"
- Pair :: "['a,'b] => 'a * 'b"
- "@Tuple" :: "args => 'a*'b" ("(1<_>)")
- Sigma :: "['a set, 'a => 'b set] => ('a*'b)set"
+ fst :: "'a * 'b => 'a"
+ snd :: "'a * 'b => 'b"
+ split :: "[['a, 'b] => 'c, 'a * 'b] => 'c"
+ prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd"
+ Pair :: "['a, 'b] => 'a * 'b"
+ Sigma :: "['a set, 'a => 'b set] => ('a * 'b) set"
- Unit :: "bool set"
- Rep_Unit :: "unit => bool"
- Abs_Unit :: "bool => unit"
- Unity :: "unit" ("<>")
+syntax
+ "@Tuple" :: "args => 'a * 'b" ("(1<_>)")
translations
+ "<x, y, z>" == "<x, <y, z>>"
+ "<x, y>" == "Pair(x, y)"
+ "<x>" => "x"
- "<x,y,z>" == "<x,<y,z>>"
- "<x,y>" == "Pair(x,y)"
- "<x>" => "x"
+defs
+ Pair_def "Pair(a, b) == Abs_Prod(Pair_Rep(a, b))"
+ fst_def "fst(p) == @a. ? b. p = <a, b>"
+ snd_def "snd(p) == @b. ? a. p = <a, b>"
+ split_def "split(c, p) == c(fst(p), snd(p))"
+ prod_fun_def "prod_fun(f, g) == split(%x y.<f(x), g(y)>)"
+ Sigma_def "Sigma(A, B) == UN x:A. UN y:B(x). {<x, y>}"
-rules
- Pair_Rep_def "Pair_Rep == (%a b. %x y. x=a & y=b)"
- Prod_def "Prod == {f. ? a b. f = Pair_Rep(a,b)}"
- (*faking a type definition...*)
- Rep_Prod "Rep_Prod(p): Prod"
- Rep_Prod_inverse "Abs_Prod(Rep_Prod(p)) = p"
- Abs_Prod_inverse "f: Prod ==> Rep_Prod(Abs_Prod(f)) = f"
- (*defining the abstract constants*)
- Pair_def "Pair(a,b) == Abs_Prod(Pair_Rep(a,b))"
- fst_def "fst(p) == @a. ? b. p = <a,b>"
- snd_def "snd(p) == @b. ? a. p = <a,b>"
- split_def "split(c,p) == c(fst(p),snd(p))"
- prod_fun_def "prod_fun(f,g) == split(%x y.<f(x), g(y)>)"
- Sigma_def "Sigma(A,B) == UN x:A. UN y:B(x). {<x,y>}"
+
+(** Unit **)
+
+subtype (Unit)
+ unit = "{p. p = True}"
- Unit_def "Unit == {p. p=True}"
- (*faking a type definition...*)
- Rep_Unit "Rep_Unit(u): Unit"
- Rep_Unit_inverse "Abs_Unit(Rep_Unit(u)) = u"
- (*defining the abstract constants*)
- Unity_def "Unity == Abs_Unit(True)"
+consts
+ Unity :: "unit" ("<>")
+
+defs
+ Unity_def "Unity == Abs_Unit(True)"
+
end
--- a/Sum.thy Fri Nov 25 13:35:32 1994 +0100
+++ b/Sum.thy Fri Nov 25 14:21:14 1994 +0100
@@ -1,55 +1,51 @@
-(* Title: HOL/sum
+(* Title: HOL/Sum.thy
ID: $Id$
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1992 University of Cambridge
-The disjoint sum of two types
+The disjoint sum of two types.
*)
Sum = Prod +
-types
- ('a,'b) "+" (infixr 10)
-
-arities
- "+" :: (term,term)term
+(* type definition *)
consts
- Inl_Rep :: "['a,'a,'b,bool] => bool"
- Inr_Rep :: "['b,'a,'b,bool] => bool"
- Sum :: "(['a,'b,bool] => bool)set"
- Rep_Sum :: "'a + 'b => (['a,'b,bool] => bool)"
- Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b"
- Inl :: "'a => 'a+'b"
- Inr :: "'b => 'a+'b"
- sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c"
+ Inl_Rep :: "['a, 'a, 'b, bool] => bool"
+ Inr_Rep :: "['b, 'a, 'b, bool] => bool"
+
+defs
+ Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)"
+ Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)"
+
+subtype (Sum)
+ ('a, 'b) "+" (infixr 10)
+ = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}"
+
+
+(* abstract constants and syntax *)
+
+consts
+ Inl :: "'a => 'a + 'b"
+ Inr :: "'b => 'a + 'b"
+ sum_case :: "['a => 'c, 'b => 'c, 'a + 'b] => 'c"
(*disjoint sum for sets; the operator + is overloaded with wrong type!*)
- "plus" :: "['a set, 'b set] => ('a+'b)set" (infixr 65)
- Part :: "['a set, 'b=>'a] => 'a set"
+ "plus" :: "['a set, 'b set] => ('a + 'b) set" (infixr 65)
+ Part :: "['a set, 'b => 'a] => 'a set"
translations
"case p of Inl(x) => a | Inr(y) => b" == "sum_case(%x.a, %y.b, p)"
-rules
- Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)"
- Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)"
-
- Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}"
- (*faking a type definition...*)
- Rep_Sum "Rep_Sum(s): Sum"
- Rep_Sum_inverse "Abs_Sum(Rep_Sum(s)) = s"
- Abs_Sum_inverse "f: Sum ==> Rep_Sum(Abs_Sum(f)) = f"
-
- (*defining the abstract constants*)
- Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
- Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
- sum_case_def "sum_case(f,g,p) == @z. (!x. p=Inl(x) --> z=f(x)) \
-\ & (!y. p=Inr(y) --> z=g(y))"
+defs
+ Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
+ Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
+ sum_case_def "sum_case(f, g, p) == @z. (!x. p=Inl(x) --> z=f(x)) \
+\ & (!y. p=Inr(y) --> z=g(y))"
sum_def "A plus B == (Inl``A) Un (Inr``B)"
(*for selecting out the components of a mutually recursive definition*)
- Part_def "Part(A,h) == A Int {x. ? z. x = h(z)}"
+ Part_def "Part(A, h) == A Int {x. ? z. x = h(z)}"
end