src/HOL/Library/Quotient_List.thy
author haftmann
Wed, 08 Dec 2010 15:05:46 +0100
changeset 41082 9ff94e7cc3b3
parent 40820 fd9c98ead9a9
child 45802 b16f976db515
child 45803 fe44c0b216ef
permissions -rw-r--r--
bot comes before top, inf before sup etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35788
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35222
diff changeset
     1
(*  Title:      HOL/Library/Quotient_List.thy
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
*)
35788
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35222
diff changeset
     4
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35222
diff changeset
     5
header {* Quotient infrastructure for the list type *}
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35222
diff changeset
     6
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
theory Quotient_List
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
imports Main Quotient_Syntax
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
begin
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
    11
declare [[map list = (map, list_all2)]]
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    13
lemma map_id [id_simps]:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    14
  "map id = id"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    15
  by (simp add: id_def fun_eq_iff map.identity)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    17
lemma list_all2_map1:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    18
  "list_all2 R (map f xs) ys \<longleftrightarrow> list_all2 (\<lambda>x. R (f x)) xs ys"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    19
  by (induct xs ys rule: list_induct2') simp_all
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    20
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    21
lemma list_all2_map2:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    22
  "list_all2 R xs (map f ys) \<longleftrightarrow> list_all2 (\<lambda>x y. R x (f y)) xs ys"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    23
  by (induct xs ys rule: list_induct2') simp_all
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    25
lemma list_all2_eq [id_simps]:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    26
  "list_all2 (op =) = (op =)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    27
proof (rule ext)+
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    28
  fix xs ys
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    29
  show "list_all2 (op =) xs ys \<longleftrightarrow> xs = ys"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    30
    by (induct xs ys rule: list_induct2') simp_all
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    31
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    33
lemma list_reflp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    34
  assumes "reflp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    35
  shows "reflp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    36
proof (rule reflpI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    37
  from assms have *: "\<And>xs. R xs xs" by (rule reflpE)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    38
  fix xs
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    39
  show "list_all2 R xs xs"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    40
    by (induct xs) (simp_all add: *)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    41
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    43
lemma list_symp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    44
  assumes "symp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    45
  shows "symp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    46
proof (rule sympI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    47
  from assms have *: "\<And>xs ys. R xs ys \<Longrightarrow> R ys xs" by (rule sympE)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    48
  fix xs ys
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    49
  assume "list_all2 R xs ys"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    50
  then show "list_all2 R ys xs"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    51
    by (induct xs ys rule: list_induct2') (simp_all add: *)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    52
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    54
lemma list_transp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    55
  assumes "transp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    56
  shows "transp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    57
proof (rule transpI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    58
  from assms have *: "\<And>xs ys zs. R xs ys \<Longrightarrow> R ys zs \<Longrightarrow> R xs zs" by (rule transpE)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    59
  fix xs ys zs
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    60
  assume A: "list_all2 R xs ys" "list_all2 R ys zs"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    61
  then have "length xs = length ys" "length ys = length zs" by (blast dest: list_all2_lengthD)+
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    62
  then show "list_all2 R xs zs" using A
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    63
    by (induct xs ys zs rule: list_induct3) (auto intro: *)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    64
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    66
lemma list_equivp [quot_equiv]:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    67
  "equivp R \<Longrightarrow> equivp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    68
  by (blast intro: equivpI list_reflp list_symp list_transp elim: equivpE)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    70
lemma list_quotient [quot_thm]:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    71
  assumes "Quotient R Abs Rep"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
    72
  shows "Quotient (list_all2 R) (map Abs) (map Rep)"
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    73
proof (rule QuotientI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    74
  from assms have "\<And>x. Abs (Rep x) = x" by (rule Quotient_abs_rep)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    75
  then show "\<And>xs. map Abs (map Rep xs) = xs" by (simp add: comp_def)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    76
next
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    77
  from assms have "\<And>x y. R (Rep x) (Rep y) \<longleftrightarrow> x = y" by (rule Quotient_rel_rep)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    78
  then show "\<And>xs. list_all2 R (map Rep xs) (map Rep xs)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    79
    by (simp add: list_all2_map1 list_all2_map2 list_all2_eq)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    80
next
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    81
  fix xs ys
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    82
  from assms have "\<And>x y. R x x \<and> R y y \<and> Abs x = Abs y \<longleftrightarrow> R x y" by (rule Quotient_rel)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    83
  then show "list_all2 R xs ys \<longleftrightarrow> list_all2 R xs xs \<and> list_all2 R ys ys \<and> map Abs xs = map Abs ys"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    84
    by (induct xs ys rule: list_induct2') auto
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    85
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    87
lemma cons_prs [quot_preserve]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
  assumes q: "Quotient R Abs Rep"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
  shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
    90
  by (auto simp add: fun_eq_iff comp_def Quotient_abs_rep [OF q])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    92
lemma cons_rsp [quot_respect]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
  assumes q: "Quotient R Abs Rep"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
    94
  shows "(R ===> list_all2 R ===> list_all2 R) (op #) (op #)"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
    95
  by auto
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    97
lemma nil_prs [quot_preserve]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
  assumes q: "Quotient R Abs Rep"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
  shows "map Abs [] = []"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
  by simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   102
lemma nil_rsp [quot_respect]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
  assumes q: "Quotient R Abs Rep"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   104
  shows "list_all2 R [] []"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
  by simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
lemma map_prs_aux:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
  shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
  by (induct l)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
     (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   114
lemma map_prs [quot_preserve]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
  shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map"
36216
8fb6cc6f3b94 respectfullness and preservation of map for identity quotients
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36154
diff changeset
   118
  and   "((abs1 ---> id) ---> map rep1 ---> id) map = map"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   119
  by (simp_all only: fun_eq_iff map_prs_aux[OF a b] comp_def)
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   120
    (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   121
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   122
lemma map_rsp [quot_respect]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
  assumes q1: "Quotient R1 Abs1 Rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
  and     q2: "Quotient R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   125
  shows "((R1 ===> R2) ===> (list_all2 R1) ===> list_all2 R2) map map"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   126
  and   "((R1 ===> op =) ===> (list_all2 R1) ===> op =) map map"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   127
  apply (simp_all add: fun_rel_def)
36216
8fb6cc6f3b94 respectfullness and preservation of map for identity quotients
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36154
diff changeset
   128
  apply(rule_tac [!] allI)+
8fb6cc6f3b94 respectfullness and preservation of map for identity quotients
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36154
diff changeset
   129
  apply(rule_tac [!] impI)
8fb6cc6f3b94 respectfullness and preservation of map for identity quotients
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36154
diff changeset
   130
  apply(rule_tac [!] allI)+
8fb6cc6f3b94 respectfullness and preservation of map for identity quotients
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36154
diff changeset
   131
  apply (induct_tac [!] xa ya rule: list_induct2')
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
  apply simp_all
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
  done
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
lemma foldr_prs_aux:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
  shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
  by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   141
lemma foldr_prs [quot_preserve]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
  shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   145
  apply (simp add: fun_eq_iff)
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   146
  by (simp only: fun_eq_iff foldr_prs_aux[OF a b])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
     (simp)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
lemma foldl_prs_aux:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
  shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
  by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   155
lemma foldl_prs [quot_preserve]:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
  assumes a: "Quotient R1 abs1 rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
  and     b: "Quotient R2 abs2 rep2"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
  shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   159
  by (simp add: fun_eq_iff foldl_prs_aux [OF a b])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   161
lemma list_all2_empty:
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   162
  shows "list_all2 R [] b \<Longrightarrow> length b = 0"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
  by (induct b) (simp_all)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
(* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
lemma foldl_rsp[quot_respect]:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
  assumes q1: "Quotient R1 Abs1 Rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
  and     q2: "Quotient R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   169
  shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_all2 R2 ===> R1) foldl foldl"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   170
  apply(auto simp add: fun_rel_def)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   171
  apply (subgoal_tac "R1 xa ya \<longrightarrow> list_all2 R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
  apply simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
  apply (rule_tac x="xa" in spec)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
  apply (rule_tac x="ya" in spec)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
  apply (rule_tac xs="xb" and ys="yb" in list_induct2)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   176
  apply (rule list_all2_lengthD)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
  apply (simp_all)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
  done
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
lemma foldr_rsp[quot_respect]:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
  assumes q1: "Quotient R1 Abs1 Rep1"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
  and     q2: "Quotient R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   183
  shows "((R1 ===> R2 ===> R2) ===> list_all2 R1 ===> R2 ===> R2) foldr foldr"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   184
  apply (auto simp add: fun_rel_def)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   185
  apply(subgoal_tac "R2 xb yb \<longrightarrow> list_all2 R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
  apply simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
  apply (rule_tac xs="xa" and ys="ya" in list_induct2)
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   188
  apply (rule list_all2_lengthD)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
  apply (simp_all)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
  done
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   192
lemma list_all2_rsp:
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   193
  assumes r: "\<forall>x y. R x y \<longrightarrow> (\<forall>a b. R a b \<longrightarrow> S x a = T y b)"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   194
  and l1: "list_all2 R x y"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   195
  and l2: "list_all2 R a b"
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   196
  shows "list_all2 S x a = list_all2 T y b"
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   197
  proof -
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   198
    have a: "length y = length x" by (rule list_all2_lengthD[OF l1, symmetric])
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   199
    have c: "length a = length b" by (rule list_all2_lengthD[OF l2])
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   200
    show ?thesis proof (cases "length x = length a")
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   201
      case True
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   202
      have b: "length x = length a" by fact
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   203
      show ?thesis using a b c r l1 l2 proof (induct rule: list_induct4)
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   204
        case Nil
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   205
        show ?case using assms by simp
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   206
      next
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   207
        case (Cons h t)
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   208
        then show ?case by auto
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   209
      qed
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   210
    next
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   211
      case False
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   212
      have d: "length x \<noteq> length a" by fact
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   213
      then have e: "\<not>list_all2 S x a" using list_all2_lengthD by auto
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   214
      have "length y \<noteq> length b" using d a c by simp
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   215
      then have "\<not>list_all2 T y b" using list_all2_lengthD by auto
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   216
      then show ?thesis using e by simp
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   217
    qed
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   218
  qed
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   219
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   220
lemma [quot_respect]:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   221
  "((R ===> R ===> op =) ===> list_all2 R ===> list_all2 R ===> op =) list_all2 list_all2"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   222
  by (simp add: list_all2_rsp fun_rel_def)
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   223
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   224
lemma [quot_preserve]:
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   225
  assumes a: "Quotient R abs1 rep1"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   226
  shows "((abs1 ---> abs1 ---> id) ---> map rep1 ---> map rep1 ---> id) list_all2 = list_all2"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   227
  apply (simp add: fun_eq_iff)
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   228
  apply clarify
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   229
  apply (induct_tac xa xb rule: list_induct2')
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   230
  apply (simp_all add: Quotient_abs_rep[OF a])
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   231
  done
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   232
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   233
lemma [quot_preserve]:
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   234
  assumes a: "Quotient R abs1 rep1"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   235
  shows "(list_all2 ((rep1 ---> rep1 ---> id) R) l m) = (l = m)"
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   236
  by (induct l m rule: list_induct2') (simp_all add: Quotient_rel_rep[OF a])
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   237
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   238
lemma list_all2_find_element:
36276
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   239
  assumes a: "x \<in> set a"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   240
  and b: "list_all2 R a b"
36276
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   241
  shows "\<exists>y. (y \<in> set b \<and> R x y)"
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   242
proof -
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   243
  have "length a = length b" using b by (rule list_all2_lengthD)
36276
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   244
  then show ?thesis using a b by (induct a b rule: list_induct2) auto
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   245
qed
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   246
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   247
lemma list_all2_refl:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   248
  assumes a: "\<And>x y. R x y = (R x = R y)"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   249
  shows "list_all2 R x x"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   250
  by (induct x) (auto simp add: a)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   251
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
end