|
1 theory QuotList |
|
2 imports QuotScript List |
|
3 begin |
|
4 |
|
5 fun |
|
6 list_rel |
|
7 where |
|
8 "list_rel R [] [] = True" |
|
9 | "list_rel R (x#xs) [] = False" |
|
10 | "list_rel R [] (x#xs) = False" |
|
11 | "list_rel R (x#xs) (y#ys) = (R x y \<and> list_rel R xs ys)" |
|
12 |
|
13 lemma list_equivp: |
|
14 assumes a: "equivp R" |
|
15 shows "equivp (list_rel R)" |
|
16 unfolding equivp_def |
|
17 apply(rule allI)+ |
|
18 apply(induct_tac x y rule: list_induct2') |
|
19 apply(simp_all add: expand_fun_eq) |
|
20 apply(metis list_rel.simps(1) list_rel.simps(2)) |
|
21 apply(metis list_rel.simps(1) list_rel.simps(2)) |
|
22 apply(rule iffI) |
|
23 apply(rule allI) |
|
24 apply(case_tac x) |
|
25 apply(simp_all) |
|
26 using a |
|
27 apply(unfold equivp_def) |
|
28 apply(auto)[1] |
|
29 apply(metis list_rel.simps(4)) |
|
30 done |
|
31 |
|
32 lemma list_rel_rel: |
|
33 assumes q: "Quotient R Abs Rep" |
|
34 shows "list_rel R r s = (list_rel R r r \<and> list_rel R s s \<and> (map Abs r = map Abs s))" |
|
35 apply(induct r s rule: list_induct2') |
|
36 apply(simp_all) |
|
37 using Quotient_rel[OF q] |
|
38 apply(metis) |
|
39 done |
|
40 |
|
41 lemma list_quotient: |
|
42 assumes q: "Quotient R Abs Rep" |
|
43 shows "Quotient (list_rel R) (map Abs) (map Rep)" |
|
44 unfolding Quotient_def |
|
45 apply(rule conjI) |
|
46 apply(rule allI) |
|
47 apply(induct_tac a) |
|
48 apply(simp) |
|
49 apply(simp add: Quotient_abs_rep[OF q]) |
|
50 apply(rule conjI) |
|
51 apply(rule allI) |
|
52 apply(induct_tac a) |
|
53 apply(simp) |
|
54 apply(simp) |
|
55 apply(simp add: Quotient_rep_reflp[OF q]) |
|
56 apply(rule allI)+ |
|
57 apply(rule list_rel_rel[OF q]) |
|
58 done |
|
59 |
|
60 |
|
61 lemma cons_prs: |
|
62 assumes q: "Quotient R Abs Rep" |
|
63 shows "(map Abs) ((Rep h) # (map Rep t)) = h # t" |
|
64 by (induct t) (simp_all add: Quotient_abs_rep[OF q]) |
|
65 |
|
66 lemma cons_rsp: |
|
67 assumes q: "Quotient R Abs Rep" |
|
68 shows "(R ===> list_rel R ===> list_rel R) op # op #" |
|
69 by (auto) |
|
70 |
|
71 lemma nil_prs: |
|
72 assumes q: "Quotient R Abs Rep" |
|
73 shows "map Abs [] \<equiv> []" |
|
74 by (simp) |
|
75 |
|
76 lemma nil_rsp: |
|
77 assumes q: "Quotient R Abs Rep" |
|
78 shows "list_rel R [] []" |
|
79 by simp |
|
80 |
|
81 lemma map_prs: |
|
82 assumes a: "Quotient R1 abs1 rep1" |
|
83 and b: "Quotient R2 abs2 rep2" |
|
84 shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l" |
|
85 by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) |
|
86 |
|
87 lemma map_rsp: |
|
88 assumes q1: "Quotient R1 Abs1 Rep1" |
|
89 and q2: "Quotient R2 Abs2 Rep2" |
|
90 shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map" |
|
91 apply(simp) |
|
92 apply(rule allI)+ |
|
93 apply(rule impI) |
|
94 apply(rule allI)+ |
|
95 apply (induct_tac xa ya rule: list_induct2') |
|
96 apply simp_all |
|
97 done |
|
98 |
|
99 (* TODO: if the above is correct, we can remove this one *) |
|
100 lemma map_rsp_lo: |
|
101 assumes q1: "Quotient R1 Abs1 Rep1" |
|
102 and q2: "Quotient R2 Abs2 Rep2" |
|
103 and a: "(R1 ===> R2) f1 f2" |
|
104 and b: "list_rel R1 l1 l2" |
|
105 shows "list_rel R2 (map f1 l1) (map f2 l2)" |
|
106 using b a |
|
107 by (induct l1 l2 rule: list_induct2') (simp_all) |
|
108 |
|
109 lemma foldr_prs: |
|
110 assumes a: "Quotient R1 abs1 rep1" |
|
111 and b: "Quotient R2 abs2 rep2" |
|
112 shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e" |
|
113 by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) |
|
114 |
|
115 lemma foldl_prs: |
|
116 assumes a: "Quotient R1 abs1 rep1" |
|
117 and b: "Quotient R2 abs2 rep2" |
|
118 shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l" |
|
119 by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) |
|
120 |
|
121 lemma list_rel_empty: "list_rel R [] b \<Longrightarrow> length b = 0" |
|
122 by (induct b) (simp_all) |
|
123 |
|
124 lemma list_rel_len: "list_rel R a b \<Longrightarrow> length a = length b" |
|
125 apply (induct a arbitrary: b) |
|
126 apply (simp add: list_rel_empty) |
|
127 apply (case_tac b) |
|
128 apply simp_all |
|
129 done |
|
130 |
|
131 (* TODO: induct_tac doesn't accept 'arbitrary'. |
|
132 induct doesn't accept 'rule'. |
|
133 that's why the proof uses manual generalisation and needs assumptions |
|
134 both in conclusion for induction and in assumptions. *) |
|
135 lemma foldl_rsp: |
|
136 assumes q1: "Quotient R1 Abs1 Rep1" |
|
137 and q2: "Quotient R2 Abs2 Rep2" |
|
138 shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl" |
|
139 apply auto |
|
140 apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)") |
|
141 apply simp |
|
142 apply (rule_tac x="xa" in spec) |
|
143 apply (rule_tac x="ya" in spec) |
|
144 apply (rule_tac xs="xb" and ys="yb" in list_induct2) |
|
145 apply (rule list_rel_len) |
|
146 apply (simp_all) |
|
147 done |
|
148 |
|
149 (* TODO: foldr_rsp should be similar *) |
|
150 |
|
151 |
|
152 |
|
153 |
|
154 (* TODO: Rest are unused *) |
|
155 |
|
156 lemma list_map_id: |
|
157 shows "map (\<lambda>x. x) = (\<lambda>x. x)" |
|
158 by simp |
|
159 |
|
160 lemma list_rel_eq: |
|
161 shows "list_rel (op =) \<equiv> (op =)" |
|
162 apply(rule eq_reflection) |
|
163 unfolding expand_fun_eq |
|
164 apply(rule allI)+ |
|
165 apply(induct_tac x xa rule: list_induct2') |
|
166 apply(simp_all) |
|
167 done |
|
168 |
|
169 lemma list_rel_refl: |
|
170 assumes a: "\<And>x y. R x y = (R x = R y)" |
|
171 shows "list_rel R x x" |
|
172 by (induct x) (auto simp add: a) |
|
173 |
|
174 end |