QuotList.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 03 Nov 2009 17:30:27 +0100
changeset 269 fe6eb116b341
parent 57 13be92f5b638
child 511 28bb34eeedc5
permissions -rw-r--r--
applic_prs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory QuotList
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports QuotScript
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
lemma LIST_map_I:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  shows "map (\<lambda>x. x) = (\<lambda>x. x)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
fun
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  LIST_REL
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
where
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  "LIST_REL R [] [] = True"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
| "LIST_REL R (x#xs) [] = False"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
| "LIST_REL R [] (x#xs) = False"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| "LIST_REL R (x#xs) (y#ys) = (R x y \<and> LIST_REL R xs ys)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
lemma LIST_REL_EQ:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  shows "LIST_REL (op =) = (op =)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
unfolding expand_fun_eq
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
apply(rule allI)+
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
apply(induct_tac x xa rule: list_induct2')
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
apply(simp_all)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
lemma LIST_REL_REFL:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  assumes a: "\<And>x y. R x y = (R x = R y)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  shows "LIST_REL R x x"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
by (induct x) (auto simp add: a)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
lemma LIST_EQUIV:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  assumes a: "EQUIV R"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  shows "EQUIV (LIST_REL R)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
unfolding EQUIV_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
apply(rule allI)+
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
apply(induct_tac x y rule: list_induct2')
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
apply(simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
apply(metis LIST_REL.simps(1) LIST_REL.simps(2))
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
apply(simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
apply(metis LIST_REL.simps(1) LIST_REL.simps(2))
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
apply(simp add: expand_fun_eq)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
apply(rule iffI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
apply(rule allI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
apply(case_tac x)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
using a
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
apply(unfold EQUIV_def)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
apply(auto)[1]
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
apply(metis LIST_REL.simps(4))
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
lemma LIST_REL_REL: 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  shows "LIST_REL R r s = (LIST_REL R r r \<and> LIST_REL R s s \<and> (map Abs r = map Abs s))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
apply(induct r s rule: list_induct2')
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
apply(simp_all)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
using QUOTIENT_REL[OF q]
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
apply(metis)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
lemma LIST_QUOTIENT:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  shows "QUOTIENT (LIST_REL R) (map Abs) (map Rep)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
unfolding QUOTIENT_def
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
apply(rule conjI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
apply(rule allI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
apply(induct_tac a)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
apply(simp add: QUOTIENT_ABS_REP[OF q])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
apply(rule conjI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
apply(rule allI)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
apply(induct_tac a)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
apply(simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
apply(simp add: QUOTIENT_REP_REFL[OF q])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
apply(rule allI)+
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
apply(rule LIST_REL_REL[OF q])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
done
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
lemma CONS_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  shows "(h#t) = (map Abs) ((Rep h)#(map Rep t))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
by (induct t) (simp_all add: QUOTIENT_ABS_REP[OF q])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
lemma CONS_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  and     a: "R h1 h2" "LIST_REL R t1 t2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  shows "LIST_REL R (h1#t1) (h2#t2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
using a by (auto)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
lemma NIL_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  shows "[] = (map Abs [])"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
by (simp)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
lemma NIL_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  assumes q: "QUOTIENT R Abs Rep"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  shows "LIST_REL R [] []"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
by simp
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
lemma MAP_PRS:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  shows "map f l = (map Abs2) (map ((Abs1 ---> Rep2) f) (map Rep1 l))"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
by (induct l)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
   (simp_all add: QUOTIENT_ABS_REP[OF q1] QUOTIENT_ABS_REP[OF q2])
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
lemma MAP_RSP:
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  assumes q1: "QUOTIENT R1 Abs1 Rep1"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  and     q2: "QUOTIENT R2 Abs2 Rep2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  and     a: "(R1 ===> R2) f1 f2" 
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  and     b: "LIST_REL R1 l1 l2"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  shows "LIST_REL R2 (map f1 l1) (map f2 l2)"
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
using b a
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
by (induct l1 l2 rule: list_induct2')
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
   (simp_all)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
end
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
(*
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
val LENGTH_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
   ("LENGTH_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
         !l. LENGTH l = LENGTH (MAP rep l)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
val LENGTH_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
   ("LENGTH_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
         !l1 l2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
          (LIST_REL R) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
          (LENGTH l1 = LENGTH l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
val APPEND_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
   ("APPEND_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
         !l m. APPEND l m = MAP abs (APPEND (MAP rep l) (MAP rep m))--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
val APPEND_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
   ("APPEND_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
         !l1 l2 m1 m2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
          (LIST_REL R) l1 l2 /\ (LIST_REL R) m1 m2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
          (LIST_REL R) (APPEND l1 m1) (APPEND l2 m2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
val FLAT_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
   ("FLAT_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
         !l. FLAT l = MAP abs (FLAT (MAP (MAP rep) l))--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
val FLAT_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
   ("FLAT_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
         !l1 l2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
          LIST_REL (LIST_REL R) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
          (LIST_REL R) (FLAT l1) (FLAT l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
val REVERSE_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
   ("REVERSE_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
         !l. REVERSE l = MAP abs (REVERSE (MAP rep l))--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
val REVERSE_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
   ("REVERSE_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
         !l1 l2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
          LIST_REL R l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
          (LIST_REL R) (REVERSE l1) (REVERSE l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
val FILTER_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
   ("FILTER_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
         !P l. FILTER P l = (MAP abs) (FILTER ((abs --> I) P) (MAP rep l))
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
       --),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
val FILTER_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
   ("FILTER_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
         !P1 P2 l1 l2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
          (R ===> $=) P1 P2 /\ (LIST_REL R) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
          (LIST_REL R) (FILTER P1 l1) (FILTER P2 l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
val NULL_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
   ("NULL_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
         !l. NULL l = NULL (MAP rep l)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
val NULL_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
   ("NULL_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
         !l1 l2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
          LIST_REL R l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
          (NULL l1 = NULL l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
val SOME_EL_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
   ("SOME_EL_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
         !l P. SOME_EL P l = SOME_EL ((abs --> I) P) (MAP rep l)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
val SOME_EL_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
   ("SOME_EL_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
         !l1 l2 P1 P2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
          (R ===> $=) P1 P2 /\ (LIST_REL R) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
          (SOME_EL P1 l1 = SOME_EL P2 l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
val ALL_EL_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
   ("ALL_EL_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
         !l P. ALL_EL P l = ALL_EL ((abs --> I) P) (MAP rep l)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
val ALL_EL_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
   ("ALL_EL_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
    (--!R (abs:'a -> 'b) rep. QUOTIENT R abs rep ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
         !l1 l2 P1 P2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
          (R ===> $=) P1 P2 /\ (LIST_REL R) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
          (ALL_EL P1 l1 = ALL_EL P2 l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
val FOLDL_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
   ("FOLDL_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
    (--!R1 (abs1:'a -> 'c) rep1. QUOTIENT R1 abs1 rep1 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
        !R2 (abs2:'b -> 'd) rep2. QUOTIENT R2 abs2 rep2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
         !l f e. FOLDL f e l =
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
                 abs1 (FOLDL ((abs1 --> abs2 --> rep1) f)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
                      (rep1 e)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
                      (MAP rep2 l))--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
val FOLDL_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
   ("FOLDL_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
    (--!R1 (abs1:'a -> 'c) rep1. QUOTIENT R1 abs1 rep1 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
        !R2 (abs2:'b -> 'd) rep2. QUOTIENT R2 abs2 rep2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
         !l1 l2 f1 f2 e1 e2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
          (R1 ===> R2 ===> R1) f1 f2 /\
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
             R1 e1 e2 /\ (LIST_REL R2) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
          R1 (FOLDL f1 e1 l1) (FOLDL f2 e2 l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
val FOLDR_PRS = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
   ("FOLDR_PRS",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
    (--!R1 (abs1:'a -> 'c) rep1. QUOTIENT R1 abs1 rep1 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
        !R2 (abs2:'b -> 'd) rep2. QUOTIENT R2 abs2 rep2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
         !l f e. FOLDR f e l =
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
                 abs2 (FOLDR ((abs1 --> abs2 --> rep2) f)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
                      (rep2 e)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
                      (MAP rep1 l))--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
val FOLDR_RSP = store_thm
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
   ("FOLDR_RSP",
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
    (--!R1 (abs1:'a -> 'c) rep1. QUOTIENT R1 abs1 rep1 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
        !R2 (abs2:'b -> 'd) rep2. QUOTIENT R2 abs2 rep2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
         !l1 l2 f1 f2 e1 e2.
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
          (R1 ===> R2 ===> R2) f1 f2 /\
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
             R2 e1 e2 /\ (LIST_REL R1) l1 l2 ==>
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
          R2 (FOLDR f1 e1 l1) (FOLDR f2 e2 l2)--),
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
*)
ebe0ea8fe247 initial commit
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254