0
|
1 |
theory Quotients
|
|
2 |
imports Main
|
|
3 |
begin
|
|
4 |
|
|
5 |
definition
|
|
6 |
"NONEMPTY E \<equiv> \<exists>x. E x x"
|
|
7 |
|
|
8 |
fun
|
|
9 |
OPTION_REL
|
|
10 |
where
|
|
11 |
"OPTION_REL R None None = True"
|
|
12 |
| "OPTION_REL R (Some x) None = False"
|
|
13 |
| "OPTION_REL R None (Some x) = False"
|
|
14 |
| "OPTION_REL R (Some x) (Some y) = R x y"
|
|
15 |
|
|
16 |
fun
|
|
17 |
option_map
|
|
18 |
where
|
|
19 |
"option_map f None = None"
|
|
20 |
| "option_map f (Some x) = Some (f x)"
|
|
21 |
|
|
22 |
fun
|
|
23 |
PROD_REL
|
|
24 |
where
|
|
25 |
"PROD_REL R1 R2 (a1,a2) (b1,b2) = (R1 a1 b1 \<and> R2 a2 b2)"
|
|
26 |
|
|
27 |
fun
|
|
28 |
prod_map
|
|
29 |
where
|
|
30 |
"prod_map f1 f2 (a,b) = (f1 a, f2 b)"
|
|
31 |
|
|
32 |
fun
|
538
|
33 |
SUM_REL
|
0
|
34 |
where
|
|
35 |
"SUM_REL R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
|
|
36 |
| "SUM_REL R1 R2 (Inl a1) (Inr b2) = False"
|
|
37 |
| "SUM_REL R1 R2 (Inr a2) (Inl b1) = False"
|
|
38 |
| "SUM_REL R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
|
|
39 |
|
|
40 |
fun
|
|
41 |
sum_map
|
|
42 |
where
|
|
43 |
"sum_map f1 f2 (Inl a) = Inl (f1 a)"
|
|
44 |
| "sum_map f1 f2 (Inr a) = Inr (f2 a)"
|
|
45 |
|
|
46 |
definition
|
|
47 |
"QUOTIENT R Abs Rep \<equiv> (\<forall>a. Abs (Rep a) = a) \<and>
|
|
48 |
(\<forall>a. R (Rep a) (Rep a)) \<and>
|
|
49 |
(\<forall>r s. R r s = (R r r \<and> R s s \<and> (Abs r = Abs s)))"
|
|
50 |
|
|
51 |
lemma QUOTIENT_PROD:
|
|
52 |
assumes a: "QUOTIENT E1 Abs1 Rep1"
|
|
53 |
and b: "QUOTIENT E2 Abs2 Rep2"
|
3
|
54 |
shows "QUOTIENT (PROD_REL E1 E2) (prod_map Abs1 Abs2) (prod_map Rep1 Rep2)"
|
0
|
55 |
using a b unfolding QUOTIENT_def
|
|
56 |
oops
|
|
57 |
|
|
58 |
lemma QUOTIENT_ABS_REP_LIST:
|
|
59 |
assumes a: "QUOTIENT_ABS_REP Abs Rep"
|
|
60 |
shows "QUOTIENT_ABS_REP (map Abs) (map Rep)"
|
|
61 |
using a
|
|
62 |
unfolding QUOTIENT_ABS_REP_def
|
|
63 |
apply(rule_tac allI)
|
|
64 |
apply(induct_tac a rule: list.induct)
|
|
65 |
apply(auto)
|
|
66 |
done
|
|
67 |
|