src/HOL/Library/Quotient_List.thy
author huffman
Sun, 22 Apr 2012 11:05:04 +0200
changeset 47660 7a5c681c0265
parent 47650 33ecf14d5ee9
child 47777 f29e7dcd7c40
permissions -rw-r--r--
new example theory for quotient/transfer
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47455
26315a545e26 updated headers;
wenzelm
parents: 47308
diff changeset
     1
(*  Title:      HOL/Library/Quotient_List.thy
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
     2
    Author:     Cezary Kaliszyk, Christian Urban and Brian Huffman
35222
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
47650
33ecf14d5ee9 add transfer rule for List.set
huffman
parents: 47649
diff changeset
     8
imports Main Quotient_Set
35222
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
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    11
subsection {* Relator for list type *}
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
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"
46663
7fe029e818c2 explicit is better than implicit
haftmann
parents: 45806
diff changeset
    15
  by (fact List.map.id)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    17
lemma list_all2_eq [id_simps, relator_eq]:
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    18
  "list_all2 (op =) = (op =)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    19
proof (rule ext)+
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    20
  fix xs ys
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    21
  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
    22
    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
    23
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
47660
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    25
lemma list_all2_OO: "list_all2 (A OO B) = list_all2 A OO list_all2 B"
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    26
proof (intro ext iffI)
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    27
  fix xs ys
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    28
  assume "list_all2 (A OO B) xs ys"
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    29
  thus "(list_all2 A OO list_all2 B) xs ys"
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    30
    unfolding OO_def
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    31
    by (induct, simp, simp add: list_all2_Cons1 list_all2_Cons2, fast)
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    32
next
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    33
  fix xs ys
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    34
  assume "(list_all2 A OO list_all2 B) xs ys"
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    35
  then obtain zs where "list_all2 A xs zs" and "list_all2 B zs ys" ..
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    36
  thus "list_all2 (A OO B) xs ys"
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    37
    by (induct arbitrary: ys, simp, clarsimp simp add: list_all2_Cons1, fast)
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    38
qed
7a5c681c0265 new example theory for quotient/transfer
huffman
parents: 47650
diff changeset
    39
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    40
lemma list_reflp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    41
  assumes "reflp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    42
  shows "reflp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    43
proof (rule reflpI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    44
  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
    45
  fix xs
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    46
  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
    47
    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
    48
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    50
lemma list_symp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    51
  assumes "symp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    52
  shows "symp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    53
proof (rule sympI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    54
  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
    55
  fix xs ys
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    56
  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
    57
  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
    58
    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
    59
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    61
lemma list_transp:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    62
  assumes "transp R"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    63
  shows "transp (list_all2 R)"
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    64
proof (rule transpI)
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    65
  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
    66
  fix xs ys zs
45803
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
    67
  assume "list_all2 R xs ys" and "list_all2 R ys zs"
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
    68
  then show "list_all2 R xs zs"
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
    69
    by (induct arbitrary: zs) (auto simp: list_all2_Cons1 intro: *)
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    70
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    72
lemma list_equivp [quot_equiv]:
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
    73
  "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
    74
  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
    75
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    76
lemma right_total_list_all2 [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    77
  "right_total R \<Longrightarrow> right_total (list_all2 R)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    78
  unfolding right_total_def
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    79
  by (rule allI, induct_tac y, simp, simp add: list_all2_Cons2)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    80
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    81
lemma right_unique_list_all2 [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    82
  "right_unique R \<Longrightarrow> right_unique (list_all2 R)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    83
  unfolding right_unique_def
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    84
  apply (rule allI, rename_tac xs, induct_tac xs)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    85
  apply (auto simp add: list_all2_Cons1)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    86
  done
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    87
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    88
lemma bi_total_list_all2 [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    89
  "bi_total A \<Longrightarrow> bi_total (list_all2 A)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    90
  unfolding bi_total_def
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    91
  apply safe
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    92
  apply (rename_tac xs, induct_tac xs, simp, simp add: list_all2_Cons1)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    93
  apply (rename_tac ys, induct_tac ys, simp, simp add: list_all2_Cons2)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    94
  done
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    95
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    96
lemma bi_unique_list_all2 [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    97
  "bi_unique A \<Longrightarrow> bi_unique (list_all2 A)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    98
  unfolding bi_unique_def
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
    99
  apply (rule conjI)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   100
  apply (rule allI, rename_tac xs, induct_tac xs)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   101
  apply (simp, force simp add: list_all2_Cons1)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   102
  apply (subst (2) all_comm, subst (1) all_comm)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   103
  apply (rule allI, rename_tac xs, induct_tac xs)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   104
  apply (simp, force simp add: list_all2_Cons2)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   105
  done
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   106
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   107
subsection {* Transfer rules for transfer package *}
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   108
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   109
lemma Nil_transfer [transfer_rule]: "(list_all2 A) [] []"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   110
  by simp
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   111
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   112
lemma Cons_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   113
  "(A ===> list_all2 A ===> list_all2 A) Cons Cons"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   114
  unfolding fun_rel_def by simp
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   115
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   116
lemma list_case_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   117
  "(B ===> (A ===> list_all2 A ===> B) ===> list_all2 A ===> B)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   118
    list_case list_case"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   119
  unfolding fun_rel_def by (simp split: list.split)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   120
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   121
lemma list_rec_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   122
  "(B ===> (A ===> list_all2 A ===> B ===> B) ===> list_all2 A ===> B)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   123
    list_rec list_rec"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   124
  unfolding fun_rel_def by (clarify, erule list_all2_induct, simp_all)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   125
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   126
lemma map_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   127
  "((A ===> B) ===> list_all2 A ===> list_all2 B) map map"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   128
  unfolding List.map_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   129
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   130
lemma append_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   131
  "(list_all2 A ===> list_all2 A ===> list_all2 A) append append"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   132
  unfolding List.append_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   133
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   134
lemma filter_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   135
  "((A ===> op =) ===> list_all2 A ===> list_all2 A) filter filter"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   136
  unfolding List.filter_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   137
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   138
lemma foldr_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   139
  "((A ===> B ===> B) ===> list_all2 A ===> B ===> B) foldr foldr"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   140
  unfolding List.foldr_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   141
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   142
lemma foldl_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   143
  "((B ===> A ===> B) ===> B ===> list_all2 A ===> B) foldl foldl"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   144
  unfolding List.foldl_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   145
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   146
lemma concat_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   147
  "(list_all2 (list_all2 A) ===> list_all2 A) concat concat"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   148
  unfolding List.concat_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   149
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   150
lemma drop_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   151
  "(op = ===> list_all2 A ===> list_all2 A) drop drop"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   152
  unfolding List.drop_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   153
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   154
lemma take_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   155
  "(op = ===> list_all2 A ===> list_all2 A) take take"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   156
  unfolding List.take_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   157
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   158
lemma length_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   159
  "(list_all2 A ===> op =) length length"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   160
  unfolding list_size_overloaded_def by transfer_prover
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   161
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   162
lemma list_all_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   163
  "((A ===> op =) ===> list_all2 A ===> op =) list_all list_all"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   164
  unfolding fun_rel_def by (clarify, erule list_all2_induct, simp_all)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   165
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   166
lemma list_all2_transfer [transfer_rule]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   167
  "((A ===> B ===> op =) ===> list_all2 A ===> list_all2 B ===> op =)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   168
    list_all2 list_all2"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   169
  apply (rule fun_relI, rule fun_relI, erule list_all2_induct)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   170
  apply (rule fun_relI, erule list_all2_induct, simp, simp)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   171
  apply (rule fun_relI, erule list_all2_induct [of B])
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   172
  apply (simp, simp add: fun_rel_def)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   173
  done
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   174
47650
33ecf14d5ee9 add transfer rule for List.set
huffman
parents: 47649
diff changeset
   175
lemma set_transfer [transfer_rule]:
33ecf14d5ee9 add transfer rule for List.set
huffman
parents: 47649
diff changeset
   176
  "(list_all2 A ===> set_rel A) set set"
33ecf14d5ee9 add transfer rule for List.set
huffman
parents: 47649
diff changeset
   177
  unfolding set_def by transfer_prover
33ecf14d5ee9 add transfer rule for List.set
huffman
parents: 47649
diff changeset
   178
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   179
subsection {* Setup for lifting package *}
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   180
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   181
lemma Quotient_list:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   182
  assumes "Quotient R Abs Rep T"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   183
  shows "Quotient (list_all2 R) (map Abs) (map Rep) (list_all2 T)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   184
proof (unfold Quotient_alt_def, intro conjI allI impI)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   185
  from assms have 1: "\<And>x y. T x y \<Longrightarrow> Abs x = y"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   186
    unfolding Quotient_alt_def by simp
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   187
  fix xs ys assume "list_all2 T xs ys" thus "map Abs xs = ys"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   188
    by (induct, simp, simp add: 1)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   189
next
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   190
  from assms have 2: "\<And>x. T (Rep x) x"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   191
    unfolding Quotient_alt_def by simp
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   192
  fix xs show "list_all2 T (map Rep xs) xs"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   193
    by (induct xs, simp, simp add: 2)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   194
next
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   195
  from assms have 3: "\<And>x y. R x y \<longleftrightarrow> T x (Abs x) \<and> T y (Abs y) \<and> Abs x = Abs y"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   196
    unfolding Quotient_alt_def by simp
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   197
  fix xs ys show "list_all2 R xs ys \<longleftrightarrow> list_all2 T xs (map Abs xs) \<and>
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   198
    list_all2 T ys (map Abs ys) \<and> map Abs xs = map Abs ys"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   199
    by (induct xs ys rule: list_induct2', simp_all, metis 3)
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   200
qed
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   201
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   202
declare [[map list = (list_all2, Quotient_list)]]
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   203
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   204
lemma list_invariant_commute [invariant_commute]:
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   205
  "list_all2 (Lifting.invariant P) = Lifting.invariant (list_all P)"
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   206
  apply (simp add: fun_eq_iff list_all2_def list_all_iff Lifting.invariant_def Ball_def) 
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   207
  apply (intro allI) 
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   208
  apply (induct_tac rule: list_induct2') 
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   209
  apply simp_all 
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   210
  apply metis
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   211
done
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   212
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   213
subsection {* Rules for quotient package *}
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   214
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   215
lemma list_quotient3 [quot_thm]:
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   216
  assumes "Quotient3 R Abs Rep"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   217
  shows "Quotient3 (list_all2 R) (map Abs) (map Rep)"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   218
proof (rule Quotient3I)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   219
  from assms have "\<And>x. Abs (Rep x) = x" by (rule Quotient3_abs_rep)
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   220
  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
   221
next
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   222
  from assms have "\<And>x y. R (Rep x) (Rep y) \<longleftrightarrow> x = y" by (rule Quotient3_rel_rep)
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   223
  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
   224
    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
   225
next
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   226
  fix xs ys
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   227
  from assms have "\<And>x y. R x x \<and> R y y \<and> Abs x = Abs y \<longleftrightarrow> R x y" by (rule Quotient3_rel)
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   228
  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
   229
    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
   230
qed
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   231
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   232
declare [[mapQ3 list = (list_all2, list_quotient3)]]
47094
1a7ad2601cb5 store the relational theorem for every relator
kuncar
parents: 46663
diff changeset
   233
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   234
lemma cons_prs [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   235
  assumes q: "Quotient3 R Abs Rep"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   236
  shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   237
  by (auto simp add: fun_eq_iff comp_def Quotient3_abs_rep [OF q])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   238
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   239
lemma cons_rsp [quot_respect]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   240
  assumes q: "Quotient3 R Abs Rep"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   241
  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
   242
  by auto
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   243
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   244
lemma nil_prs [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   245
  assumes q: "Quotient3 R Abs Rep"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   246
  shows "map Abs [] = []"
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   247
  by simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   248
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   249
lemma nil_rsp [quot_respect]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   250
  assumes q: "Quotient3 R Abs Rep"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   251
  shows "list_all2 R [] []"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   252
  by simp
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   253
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   254
lemma map_prs_aux:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   255
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   256
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   257
  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
   258
  by (induct l)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   259
     (simp_all add: Quotient3_abs_rep[OF a] Quotient3_abs_rep[OF b])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   260
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   261
lemma map_prs [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   262
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   263
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   264
  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
   265
  and   "((abs1 ---> id) ---> map rep1 ---> id) map = map"
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   266
  by (simp_all only: fun_eq_iff map_prs_aux[OF a b] comp_def)
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   267
    (simp_all add: Quotient3_abs_rep[OF a] Quotient3_abs_rep[OF b])
40463
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   268
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   269
lemma map_rsp [quot_respect]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   270
  assumes q1: "Quotient3 R1 Abs1 Rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   271
  and     q2: "Quotient3 R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   272
  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
   273
  and   "((R1 ===> op =) ===> (list_all2 R1) ===> op =) map map"
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   274
  unfolding list_all2_eq [symmetric] by (rule map_transfer)+
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   275
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   276
lemma foldr_prs_aux:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   277
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   278
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   279
  shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   280
  by (induct l) (simp_all add: Quotient3_abs_rep[OF a] Quotient3_abs_rep[OF b])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   281
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   282
lemma foldr_prs [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   283
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   284
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   285
  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
   286
  apply (simp add: fun_eq_iff)
75e544159549 fun_rel_def is no simp rule by default
haftmann
parents: 40032
diff changeset
   287
  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
   288
     (simp)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   289
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   290
lemma foldl_prs_aux:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   291
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   292
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   293
  shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   294
  by (induct l arbitrary:e) (simp_all add: Quotient3_abs_rep[OF a] Quotient3_abs_rep[OF b])
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   295
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   296
lemma foldl_prs [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   297
  assumes a: "Quotient3 R1 abs1 rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   298
  and     b: "Quotient3 R2 abs2 rep2"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   299
  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
   300
  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
   301
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   302
(* 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
   303
lemma foldl_rsp[quot_respect]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   304
  assumes q1: "Quotient3 R1 Abs1 Rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   305
  and     q2: "Quotient3 R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   306
  shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_all2 R2 ===> R1) foldl foldl"
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   307
  by (rule foldl_transfer)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   308
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   309
lemma foldr_rsp[quot_respect]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   310
  assumes q1: "Quotient3 R1 Abs1 Rep1"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   311
  and     q2: "Quotient3 R2 Abs2 Rep2"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   312
  shows "((R1 ===> R2 ===> R2) ===> list_all2 R1 ===> R2 ===> R2) foldr foldr"
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   313
  by (rule foldr_transfer)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   314
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   315
lemma list_all2_rsp:
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   316
  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
   317
  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
   318
  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
   319
  shows "list_all2 S x a = list_all2 T y b"
45803
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
   320
  using l1 l2
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
   321
  by (induct arbitrary: a b rule: list_all2_induct,
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
   322
    auto simp: list_all2_Cons1 list_all2_Cons2 r)
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   323
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   324
lemma [quot_respect]:
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   325
  "((R ===> R ===> op =) ===> list_all2 R ===> list_all2 R ===> op =) list_all2 list_all2"
47641
2cddc27a881f new transfer package rules and lifting setup for lists
huffman
parents: 47634
diff changeset
   326
  by (rule list_all2_transfer)
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   327
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   328
lemma [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   329
  assumes a: "Quotient3 R abs1 rep1"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   330
  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
   331
  apply (simp add: fun_eq_iff)
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   332
  apply clarify
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   333
  apply (induct_tac xa xb rule: list_induct2')
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   334
  apply (simp_all add: Quotient3_abs_rep[OF a])
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   335
  done
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   336
40820
fd9c98ead9a9 more systematic and compact proofs on type relation operators using natural deduction rules
haftmann
parents: 40463
diff changeset
   337
lemma [quot_preserve]:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   338
  assumes a: "Quotient3 R abs1 rep1"
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   339
  shows "(list_all2 ((rep1 ---> rep1 ---> id) R) l m) = (l = m)"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 47094
diff changeset
   340
  by (induct l m rule: list_induct2') (simp_all add: Quotient3_rel_rep[OF a])
36154
11c6106d7787 Respectfullness and preservation of list_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35788
diff changeset
   341
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   342
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
   343
  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
   344
  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
   345
  shows "\<exists>y. (y \<in> set b \<and> R x y)"
45803
fe44c0b216ef remove some duplicate lemmas, simplify some proofs
huffman
parents: 40820
diff changeset
   346
  using b a by induct auto
36276
92011cc923f5 fun_rel introduction and list_rel elimination for quotient package
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36216
diff changeset
   347
37492
ab36b1a50ca8 Replace 'list_rel' by 'list_all2'; they are equivalent.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 36812
diff changeset
   348
lemma list_all2_refl:
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   349
  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
   350
  shows "list_all2 R x x"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   351
  by (induct x) (auto simp add: a)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   352
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   353
end