author | lcp |
Thu, 18 Aug 1994 11:43:40 +0200 | |
changeset 107 | 960e332d2e70 |
parent 51 | 934a58983311 |
child 117 | 3716c99fb6a1 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/sum |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
The disjoint sum of two types |
|
7 |
*) |
|
8 |
||
9 |
Sum = Prod + |
|
51 | 10 |
|
11 |
types |
|
12 |
('a,'b) "+" (infixl 10) |
|
13 |
||
14 |
arities |
|
15 |
"+" :: (term,term)term |
|
16 |
||
0 | 17 |
consts |
51 | 18 |
Inl_Rep :: "['a,'a,'b,bool] => bool" |
19 |
Inr_Rep :: "['b,'a,'b,bool] => bool" |
|
20 |
Sum :: "(['a,'b,bool] => bool)set" |
|
21 |
Rep_Sum :: "'a + 'b => (['a,'b,bool] => bool)" |
|
22 |
Abs_Sum :: "(['a,'b,bool] => bool) => 'a+'b" |
|
23 |
Inl :: "'a => 'a+'b" |
|
24 |
Inr :: "'b => 'a+'b" |
|
107
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
25 |
sum_case :: "['a=>'c,'b=>'c, 'a+'b] =>'c" |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
26 |
Part :: "['a set, 'a=>'a] => 'a set" |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
27 |
|
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
28 |
translations |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
29 |
"case p of Inl(x) => a | Inr(y) => b" == "sum_case(%x.a, %y.b, p)" |
51 | 30 |
|
0 | 31 |
rules |
32 |
Inl_Rep_def "Inl_Rep == (%a. %x y p. x=a & p)" |
|
33 |
Inr_Rep_def "Inr_Rep == (%b. %x y p. y=b & ~p)" |
|
107
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
34 |
|
0 | 35 |
Sum_def "Sum == {f. (? a. f = Inl_Rep(a)) | (? b. f = Inr_Rep(b))}" |
36 |
(*faking a type definition...*) |
|
37 |
Rep_Sum "Rep_Sum(s): Sum" |
|
38 |
Rep_Sum_inverse "Abs_Sum(Rep_Sum(s)) = s" |
|
39 |
Abs_Sum_inverse "f: Sum ==> Rep_Sum(Abs_Sum(f)) = f" |
|
107
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
40 |
|
0 | 41 |
(*defining the abstract constants*) |
42 |
Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" |
|
43 |
Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" |
|
107
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
44 |
sum_case_def "sum_case(f,g,p) == @z. (!x. p=Inl(x) --> z=f(x)) \ |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
45 |
\ & (!y. p=Inr(y) --> z=g(y))" |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
46 |
|
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
47 |
(*for selecting out the components of a mutually recursive definition*) |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
48 |
Part_def "Part(A,h) == A Int {x. ? z. x = h(z)}" |
960e332d2e70
HOL/Sum: rotated arguments of sum_case; added translation for case macro
lcp
parents:
51
diff
changeset
|
49 |
|
0 | 50 |
end |