| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Tue, 26 Jan 2010 08:55:55 +0100 | |
| changeset 929 | e812f58fd128 | 
| parent 927 | 04ef4bd3b936 | 
| child 936 | da5e4b8317c7 | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
theory QuotList  | 
| 
600
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
2  | 
imports QuotMain List  | 
| 0 | 3  | 
begin  | 
4  | 
||
5  | 
fun  | 
|
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
6  | 
list_rel  | 
| 0 | 7  | 
where  | 
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
8  | 
"list_rel R [] [] = True"  | 
| 
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
9  | 
| "list_rel R (x#xs) [] = False"  | 
| 
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
10  | 
| "list_rel R [] (x#xs) = False"  | 
| 
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
11  | 
| "list_rel R (x#xs) (y#ys) = (R x y \<and> list_rel R xs ys)"  | 
| 0 | 12  | 
|
| 
600
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
13  | 
declare [[map list = (map, list_rel)]]  | 
| 
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
14  | 
|
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
15  | 
lemma list_equivp[quot_equiv]:  | 
| 529 | 16  | 
assumes a: "equivp R"  | 
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
17  | 
shows "equivp (list_rel R)"  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
18  | 
unfolding equivp_def  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
19  | 
apply(rule allI)+  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
20  | 
apply(induct_tac x y rule: list_induct2')  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
21  | 
apply(simp_all add: expand_fun_eq)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
22  | 
apply(metis list_rel.simps(1) list_rel.simps(2))  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
23  | 
apply(metis list_rel.simps(1) list_rel.simps(2))  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
24  | 
apply(rule iffI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
25  | 
apply(rule allI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
26  | 
apply(case_tac x)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
27  | 
apply(simp_all)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
28  | 
using a  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
29  | 
apply(unfold equivp_def)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
30  | 
apply(auto)[1]  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
31  | 
apply(metis list_rel.simps(4))  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
32  | 
done  | 
| 0 | 33  | 
|
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
34  | 
lemma list_rel_rel:  | 
| 529 | 35  | 
assumes q: "Quotient R Abs Rep"  | 
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
36  | 
shows "list_rel R r s = (list_rel R r r \<and> list_rel R s s \<and> (map Abs r = map Abs s))"  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
37  | 
apply(induct r s rule: list_induct2')  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
38  | 
apply(simp_all)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
39  | 
using Quotient_rel[OF q]  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
40  | 
apply(metis)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
41  | 
done  | 
| 0 | 42  | 
|
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
43  | 
lemma list_quotient[quot_thm]:  | 
| 529 | 44  | 
assumes q: "Quotient R Abs Rep"  | 
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
45  | 
shows "Quotient (list_rel R) (map Abs) (map Rep)"  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
46  | 
unfolding Quotient_def  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
47  | 
apply(rule conjI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
48  | 
apply(rule allI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
49  | 
apply(induct_tac a)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
50  | 
apply(simp)  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
51  | 
apply(simp add: Quotient_abs_rep[OF q])  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
52  | 
apply(rule conjI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
53  | 
apply(rule allI)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
54  | 
apply(induct_tac a)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
55  | 
apply(simp)  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
56  | 
apply(simp)  | 
| 
541
 
94deffed619d
more name cleaning and removing
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
540 
diff
changeset
 | 
57  | 
apply(simp add: Quotient_rep_reflp[OF q])  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
58  | 
apply(rule allI)+  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
59  | 
apply(rule list_rel_rel[OF q])  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
60  | 
done  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
61  | 
|
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
62  | 
lemma map_id[id_simps]: "map id = id"  | 
| 
600
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
63  | 
apply (rule ext)  | 
| 
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
64  | 
apply (rule_tac list="x" in list.induct)  | 
| 
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
65  | 
apply (simp_all)  | 
| 
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
66  | 
done  | 
| 
 
5d932e7a856c
List moved after QuotMain
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
597 
diff
changeset
 | 
67  | 
|
| 
644
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
68  | 
lemma cons_prs_aux:  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
69  | 
assumes q: "Quotient R Abs Rep"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
70  | 
shows "(map Abs) ((Rep h) # (map Rep t)) = h # t"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
71  | 
by (induct t) (simp_all add: Quotient_abs_rep[OF q])  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
72  | 
|
| 
644
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
73  | 
lemma cons_prs[quot_preserve]:  | 
| 
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
74  | 
assumes q: "Quotient R Abs Rep"  | 
| 
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
75  | 
shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
76  | 
by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q]) (simp)  | 
| 
644
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
77  | 
|
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
78  | 
lemma cons_rsp[quot_respect]:  | 
| 529 | 79  | 
assumes q: "Quotient R Abs Rep"  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
80  | 
shows "(R ===> list_rel R ===> list_rel R) op # op #"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
81  | 
by auto  | 
| 0 | 82  | 
|
| 
644
 
97a397ba5743
started to reformulate preserve lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
636 
diff
changeset
 | 
83  | 
lemma nil_prs[quot_preserve]:  | 
| 529 | 84  | 
assumes q: "Quotient R Abs Rep"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
85  | 
shows "map Abs [] = []"  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
86  | 
by simp  | 
| 0 | 87  | 
|
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
88  | 
lemma nil_rsp[quot_respect]:  | 
| 529 | 89  | 
assumes q: "Quotient R Abs Rep"  | 
| 
537
 
57073b0b8fac
Even more name changes and cleaning
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
533 
diff
changeset
 | 
90  | 
shows "list_rel R [] []"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
91  | 
by simp  | 
| 0 | 92  | 
|
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
93  | 
lemma map_prs_aux:  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
94  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
95  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
96  | 
shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
97  | 
by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])  | 
| 0 | 98  | 
|
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
99  | 
|
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
100  | 
lemma map_prs[quot_preserve]:  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
101  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
102  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
103  | 
shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
104  | 
by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b]) (simp)  | 
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
105  | 
|
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
106  | 
|
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
107  | 
lemma map_rsp[quot_respect]:  | 
| 529 | 108  | 
assumes q1: "Quotient R1 Abs1 Rep1"  | 
109  | 
and q2: "Quotient R2 Abs2 Rep2"  | 
|
| 
565
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
110  | 
shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
111  | 
apply(simp)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
112  | 
apply(rule allI)+  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
113  | 
apply(rule impI)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
114  | 
apply(rule allI)+  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
115  | 
apply (induct_tac xa ya rule: list_induct2')  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
116  | 
apply simp_all  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
117  | 
done  | 
| 
565
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
118  | 
|
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
119  | 
lemma foldr_prs_aux:  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
120  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
121  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
122  | 
shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
123  | 
by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
124  | 
|
| 
796
 
64f9c76f70c7
corrected wrong [quot_respect] attribute; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
768 
diff
changeset
 | 
125  | 
lemma foldr_prs[quot_preserve]:  | 
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
126  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
127  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
128  | 
shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
129  | 
by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b]) (simp)  | 
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
130  | 
|
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
131  | 
lemma foldl_prs_aux:  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
132  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
133  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
134  | 
shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
135  | 
by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])  | 
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
136  | 
|
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
137  | 
lemma foldl_prs[quot_preserve]:  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
138  | 
assumes a: "Quotient R1 abs1 rep1"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
139  | 
and b: "Quotient R2 abs2 rep2"  | 
| 
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
140  | 
shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
141  | 
by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b]) (simp)  | 
| 
645
 
fe2a37cfecd3
proper formulation of all preservation theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
644 
diff
changeset
 | 
142  | 
|
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
143  | 
lemma list_rel_empty:  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
144  | 
"list_rel R [] b \<Longrightarrow> length b = 0"  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
145  | 
by (induct b) (simp_all)  | 
| 0 | 146  | 
|
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
147  | 
lemma list_rel_len:  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
148  | 
"list_rel R a b \<Longrightarrow> length a = length b"  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
149  | 
apply (induct a arbitrary: b)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
150  | 
apply (simp add: list_rel_empty)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
151  | 
apply (case_tac b)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
152  | 
apply simp_all  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
153  | 
done  | 
| 0 | 154  | 
|
| 668 | 155  | 
(* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *)  | 
| 
636
 
520a4084d064
changed names of attributes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
623 
diff
changeset
 | 
156  | 
lemma foldl_rsp[quot_respect]:  | 
| 
565
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
157  | 
assumes q1: "Quotient R1 Abs1 Rep1"  | 
| 
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
158  | 
and q2: "Quotient R2 Abs2 Rep2"  | 
| 
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
159  | 
shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
160  | 
apply(auto)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
161  | 
apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
162  | 
apply simp  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
163  | 
apply (rule_tac x="xa" in spec)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
164  | 
apply (rule_tac x="ya" in spec)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
165  | 
apply (rule_tac xs="xb" and ys="yb" in list_induct2)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
166  | 
apply (rule list_rel_len)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
167  | 
apply (simp_all)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
168  | 
done  | 
| 
565
 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
541 
diff
changeset
 | 
169  | 
|
| 666 | 170  | 
lemma foldr_rsp[quot_respect]:  | 
171  | 
assumes q1: "Quotient R1 Abs1 Rep1"  | 
|
172  | 
and q2: "Quotient R2 Abs2 Rep2"  | 
|
173  | 
shows "((R1 ===> R2 ===> R2) ===> list_rel R1 ===> R2 ===> R2) foldr foldr"  | 
|
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
174  | 
apply auto  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
175  | 
apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
176  | 
apply simp  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
177  | 
apply (rule_tac xs="xa" and ys="ya" in list_induct2)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
178  | 
apply (rule list_rel_len)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
179  | 
apply (simp_all)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
180  | 
done  | 
| 
540
 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
539 
diff
changeset
 | 
181  | 
|
| 
825
 
970e86082cd7
Modifictaions for new_relation.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
796 
diff
changeset
 | 
182  | 
lemma list_rel_eq[id_simps]:  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
183  | 
shows "(list_rel (op =)) = (op =)"  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
184  | 
unfolding expand_fun_eq  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
185  | 
apply(rule allI)+  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
186  | 
apply(induct_tac x xa rule: list_induct2')  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
187  | 
apply(simp_all)  | 
| 
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
188  | 
done  | 
| 0 | 189  | 
|
| 
541
 
94deffed619d
more name cleaning and removing
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
540 
diff
changeset
 | 
190  | 
lemma list_rel_refl:  | 
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
191  | 
assumes a: "\<And>x y. R x y = (R x = R y)"  | 
| 
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
192  | 
shows "list_rel R x x"  | 
| 
927
 
04ef4bd3b936
more eq_reflection & other cleaning.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
924 
diff
changeset
 | 
193  | 
by (induct x) (auto simp add: a)  | 
| 0 | 194  | 
|
| 
539
 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
537 
diff
changeset
 | 
195  | 
end  |