2014-12-23 22:34:16 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Parikshit Khanna. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
|
|
|
|
|
Module: data.list.basic
|
|
|
|
|
Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura
|
|
|
|
|
|
|
|
|
|
Basic properties of lists.
|
|
|
|
|
-/
|
2014-12-01 04:34:12 +00:00
|
|
|
|
import logic tools.helper_tactics data.nat.basic
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-05 04:30:19 +00:00
|
|
|
|
open eq.ops helper_tactics nat prod function
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
inductive list (T : Type) : Type :=
|
2015-02-26 01:00:10 +00:00
|
|
|
|
| nil {} : list T
|
|
|
|
|
| cons : T → list T → list T
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
namespace list
|
2014-10-21 21:08:07 +00:00
|
|
|
|
notation h :: t := cons h t
|
2015-03-14 05:25:21 +00:00
|
|
|
|
notation `[` l:(foldr `,` (h t, cons h t) nil `]`) := l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-09 14:13:06 +00:00
|
|
|
|
variable {T : Type}
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- append -/
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition append : list T → list T → list T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] l := l
|
|
|
|
|
| (h :: s) t := h :: (append s t)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-21 21:08:07 +00:00
|
|
|
|
notation l₁ ++ l₂ := append l₁ l₂
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem append_nil_left (t : list T) : [] ++ t = t
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem append_cons (x : T) (s t : list T) : (x::s) ++ t = x::(s ++ t)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem append_nil_right : ∀ (t : list T), t ++ [] = t
|
|
|
|
|
| [] := rfl
|
|
|
|
|
| (a :: l) := calc
|
|
|
|
|
(a :: l) ++ [] = a :: (l ++ []) : rfl
|
|
|
|
|
... = a :: l : append_nil_right l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-09 02:47:44 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
theorem append.assoc : ∀ (s t u : list T), s ++ t ++ u = s ++ (t ++ u)
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] t u := rfl
|
|
|
|
|
| (a :: l) t u :=
|
2015-03-05 04:40:06 +00:00
|
|
|
|
show a :: (l ++ t ++ u) = (a :: l) ++ (t ++ u),
|
|
|
|
|
by rewrite (append.assoc l t u)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- length -/
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition length : list T → nat
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := 0
|
|
|
|
|
| (a :: l) := length l + 1
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem length_nil : length (@nil T) = 0
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem length_cons (x : T) (t : list T) : length (x::t) = length t + 1
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
theorem length_append : ∀ (s t : list T), length (s ++ t) = length s + length t
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] t := calc
|
|
|
|
|
length ([] ++ t) = length t : rfl
|
|
|
|
|
... = length [] + length t : zero_add
|
|
|
|
|
| (a :: s) t := calc
|
2015-02-26 00:20:44 +00:00
|
|
|
|
length (a :: s ++ t) = length (s ++ t) + 1 : rfl
|
|
|
|
|
... = length s + length t + 1 : length_append
|
|
|
|
|
... = (length s + 1) + length t : add.succ_left
|
|
|
|
|
... = length (a :: s) + length t : rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- add_rewrite length_nil length_cons
|
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- concat -/
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition concat : Π (x : T), list T → list T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| a [] := [a]
|
|
|
|
|
| a (b :: l) := b :: concat a l
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem concat_nil (x : T) : concat x [] = [x]
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem concat_cons (x y : T) (l : list T) : concat x (y::l) = y::(concat x l)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
theorem concat_eq_append (a : T) : ∀ (l : list T), concat a l = l ++ [a]
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := rfl
|
|
|
|
|
| (b :: l) :=
|
|
|
|
|
show b :: (concat a l) = (b :: l) ++ (a :: []),
|
2015-03-05 04:40:06 +00:00
|
|
|
|
by rewrite concat_eq_append
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
|
|
|
|
-- add_rewrite append_nil append_cons
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- reverse -/
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition reverse : list T → list T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := []
|
|
|
|
|
| (a :: l) := concat a (reverse l)
|
2015-01-07 21:38:11 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem reverse_nil : reverse (@nil T) = []
|
2015-01-07 21:38:11 +00:00
|
|
|
|
|
|
|
|
|
theorem reverse_cons (x : T) (l : list T) : reverse (x::l) = concat x (reverse l)
|
|
|
|
|
|
|
|
|
|
theorem reverse_singleton (x : T) : reverse [x] = [x]
|
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
theorem reverse_append : ∀ (s t : list T), reverse (s ++ t) = (reverse t) ++ (reverse s)
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] t2 := calc
|
|
|
|
|
reverse ([] ++ t2) = reverse t2 : rfl
|
|
|
|
|
... = (reverse t2) ++ [] : append_nil_right
|
|
|
|
|
... = (reverse t2) ++ (reverse []) : by rewrite reverse_nil
|
|
|
|
|
| (a2 :: s2) t2 := calc
|
2015-02-26 00:20:44 +00:00
|
|
|
|
reverse ((a2 :: s2) ++ t2) = concat a2 (reverse (s2 ++ t2)) : rfl
|
|
|
|
|
... = concat a2 (reverse t2 ++ reverse s2) : reverse_append
|
|
|
|
|
... = (reverse t2 ++ reverse s2) ++ [a2] : concat_eq_append
|
|
|
|
|
... = reverse t2 ++ (reverse s2 ++ [a2]) : append.assoc
|
|
|
|
|
... = reverse t2 ++ concat a2 (reverse s2) : concat_eq_append
|
|
|
|
|
... = reverse t2 ++ reverse (a2 :: s2) : rfl
|
|
|
|
|
|
|
|
|
|
theorem reverse_reverse : ∀ (l : list T), reverse (reverse l) = l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := rfl
|
|
|
|
|
| (a :: l) := calc
|
2015-02-26 00:20:44 +00:00
|
|
|
|
reverse (reverse (a :: l)) = reverse (concat a (reverse l)) : rfl
|
|
|
|
|
... = reverse (reverse l ++ [a]) : concat_eq_append
|
|
|
|
|
... = reverse [a] ++ reverse (reverse l) : reverse_append
|
|
|
|
|
... = reverse [a] ++ l : reverse_reverse
|
|
|
|
|
... = a :: l : rfl
|
2015-01-07 21:38:11 +00:00
|
|
|
|
|
|
|
|
|
theorem concat_eq_reverse_cons (x : T) (l : list T) : concat x l = reverse (x :: reverse l) :=
|
|
|
|
|
calc
|
2015-01-09 02:47:44 +00:00
|
|
|
|
concat x l = concat x (reverse (reverse l)) : reverse_reverse
|
2015-01-07 21:38:11 +00:00
|
|
|
|
... = reverse (x :: reverse l) : rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- head and tail -/
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition head [h : inhabited T] : list T → T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := arbitrary T
|
|
|
|
|
| (a :: l) := a
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem head_cons [h : inhabited T] (a : T) (l : list T) : head (a::l) = a
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
theorem head_append [h : inhabited T] (t : list T) : ∀ {s : list T}, s ≠ [] → head (s ++ t) = head s
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] H := absurd rfl H
|
|
|
|
|
| (a :: s) H :=
|
2015-03-05 04:40:06 +00:00
|
|
|
|
show head (a :: (s ++ t)) = head (a :: s),
|
|
|
|
|
by rewrite head_cons
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition tail : list T → list T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := []
|
|
|
|
|
| (a :: l) := l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem tail_nil : tail (@nil T) = []
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem tail_cons (a : T) (l : list T) : tail (a::l) = l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem cons_head_tail [h : inhabited T] {l : list T} : l ≠ [] → (head l)::(tail l) = l :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.cases_on l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
(assume H : [] ≠ [], absurd rfl H)
|
|
|
|
|
(take x l, assume H : x::l ≠ [], rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- list membership -/
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition mem : T → list T → Prop
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| a [] := false
|
|
|
|
|
| a (b :: l) := a = b ∨ mem a l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-21 21:08:07 +00:00
|
|
|
|
notation e ∈ s := mem e s
|
2015-03-31 02:11:45 +00:00
|
|
|
|
notation e ∉ s := ¬ e ∈ s
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem mem_nil (x : T) : x ∈ [] ↔ false :=
|
2014-09-10 14:55:32 +00:00
|
|
|
|
iff.rfl
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-03-31 02:11:45 +00:00
|
|
|
|
theorem not_mem_nil (x : T) : x ∉ [] :=
|
|
|
|
|
iff.mp !mem_nil
|
|
|
|
|
|
2015-03-29 20:59:36 +00:00
|
|
|
|
theorem mem_cons (x : T) (l : list T) : x ∈ x :: l :=
|
|
|
|
|
or.inl rfl
|
|
|
|
|
|
2015-03-31 02:11:45 +00:00
|
|
|
|
theorem mem_cons_of_mem (x : T) {y : T} {l : list T} : x ∈ l → x ∈ y :: l :=
|
|
|
|
|
assume H, or.inr H
|
|
|
|
|
|
2015-03-29 20:59:36 +00:00
|
|
|
|
theorem mem_cons_iff (x y : T) (l : list T) : x ∈ y::l ↔ (x = y ∨ x ∈ l) :=
|
2014-09-10 14:55:32 +00:00
|
|
|
|
iff.rfl
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
theorem mem_or_mem_of_mem_append {x : T} {s t : list T} : x ∈ s ++ t → x ∈ s ∨ x ∈ t :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.induction_on s or.inr
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(take y s,
|
|
|
|
|
assume IH : x ∈ s ++ t → x ∈ s ∨ x ∈ t,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
assume H1 : x ∈ y::s ++ t,
|
2014-07-31 20:33:35 +00:00
|
|
|
|
have H2 : x = y ∨ x ∈ s ++ t, from H1,
|
2014-12-15 20:05:44 +00:00
|
|
|
|
have H3 : x = y ∨ x ∈ s ∨ x ∈ t, from or_of_or_of_imp_right H2 IH,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.elim_right or.assoc H3)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
theorem mem_append_of_mem_or_mem {x : T} {s t : list T} : x ∈ s ∨ x ∈ t → x ∈ s ++ t :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.induction_on s
|
2014-12-15 20:05:44 +00:00
|
|
|
|
(take H, or.elim H false.elim (assume H, H))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(take y s,
|
|
|
|
|
assume IH : x ∈ s ∨ x ∈ t → x ∈ s ++ t,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
assume H : x ∈ y::s ∨ x ∈ t,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim H
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(assume H1,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim H1
|
|
|
|
|
(take H2 : x = y, or.inl H2)
|
|
|
|
|
(take H2 : x ∈ s, or.inr (IH (or.inl H2))))
|
|
|
|
|
(assume H1 : x ∈ t, or.inr (IH (or.inr H1))))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
theorem mem_append_iff (x : T) (s t : list T) : x ∈ s ++ t ↔ x ∈ s ∨ x ∈ t :=
|
|
|
|
|
iff.intro mem_or_mem_of_mem_append mem_append_of_mem_or_mem
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-01-26 19:31:12 +00:00
|
|
|
|
local attribute mem [reducible]
|
|
|
|
|
local attribute append [reducible]
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem mem_split {x : T} {l : list T} : x ∈ l → ∃s t : list T, l = s ++ (x::t) :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.induction_on l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
(take H : x ∈ [], false.elim (iff.elim_left !mem_nil H))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(take y l,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
assume IH : x ∈ l → ∃s t : list T, l = s ++ (x::t),
|
|
|
|
|
assume H : x ∈ y::l,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
or.elim H
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(assume H1 : x = y,
|
2015-03-14 05:25:21 +00:00
|
|
|
|
exists.intro [] (!exists.intro (H1 ▸ rfl)))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(assume H1 : x ∈ l,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
obtain s (H2 : ∃t : list T, l = s ++ (x::t)), from IH H1,
|
|
|
|
|
obtain t (H3 : l = s ++ (x::t)), from H2,
|
|
|
|
|
have H4 : y :: l = (y::s) ++ (x::t),
|
2014-09-05 01:41:06 +00:00
|
|
|
|
from H3 ▸ rfl,
|
2014-12-16 03:05:03 +00:00
|
|
|
|
!exists.intro (!exists.intro H4)))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-04-01 01:38:05 +00:00
|
|
|
|
theorem mem_append_left {a : T} {l₁ : list T} (l₂ : list T) : a ∈ l₁ → a ∈ l₁ ++ l₂ :=
|
|
|
|
|
assume ainl₁, mem_append_of_mem_or_mem (or.inl ainl₁)
|
|
|
|
|
|
|
|
|
|
theorem mem_append_right {a : T} (l₁ : list T) {l₂ : list T} : a ∈ l₂ → a ∈ l₁ ++ l₂ :=
|
|
|
|
|
assume ainl₂, mem_append_of_mem_or_mem (or.inr ainl₂)
|
|
|
|
|
|
2015-02-25 20:34:49 +00:00
|
|
|
|
definition decidable_mem [instance] [H : decidable_eq T] (x : T) (l : list T) : decidable (x ∈ l) :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.rec_on l
|
2015-01-07 21:38:11 +00:00
|
|
|
|
(decidable.inr (not_of_iff_false !mem_nil))
|
2014-10-05 20:09:56 +00:00
|
|
|
|
(take (h : T) (l : list T) (iH : decidable (x ∈ l)),
|
2014-09-10 23:42:27 +00:00
|
|
|
|
show decidable (x ∈ h::l), from
|
2014-09-08 04:06:32 +00:00
|
|
|
|
decidable.rec_on iH
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(assume Hp : x ∈ l,
|
2014-09-08 04:06:32 +00:00
|
|
|
|
decidable.rec_on (H x h)
|
|
|
|
|
(assume Heq : x = h,
|
|
|
|
|
decidable.inl (or.inl Heq))
|
|
|
|
|
(assume Hne : x ≠ h,
|
|
|
|
|
decidable.inl (or.inr Hp)))
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(assume Hn : ¬x ∈ l,
|
2014-09-08 04:06:32 +00:00
|
|
|
|
decidable.rec_on (H x h)
|
|
|
|
|
(assume Heq : x = h,
|
|
|
|
|
decidable.inl (or.inl Heq))
|
|
|
|
|
(assume Hne : x ≠ h,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
have H1 : ¬(x = h ∨ x ∈ l), from
|
|
|
|
|
assume H2 : x = h ∨ x ∈ l, or.elim H2
|
2014-09-08 04:06:32 +00:00
|
|
|
|
(assume Heq, absurd Heq Hne)
|
|
|
|
|
(assume Hp, absurd Hp Hn),
|
2014-09-10 23:42:27 +00:00
|
|
|
|
have H2 : ¬x ∈ h::l, from
|
2015-03-29 20:59:36 +00:00
|
|
|
|
iff.elim_right (not_iff_not_of_iff !mem_cons_iff) H1,
|
2014-09-08 04:06:32 +00:00
|
|
|
|
decidable.inr H2)))
|
|
|
|
|
|
2015-03-30 10:18:02 +00:00
|
|
|
|
theorem mem_of_ne_of_mem {x y : T} {l : list T} (H₁ : x ≠ y) (H₂ : x ∈ y :: l) : x ∈ l :=
|
|
|
|
|
or.elim H₂ (λe, absurd e H₁) (λr, r)
|
|
|
|
|
|
|
|
|
|
definition sublist (l₁ l₂ : list T) := ∀ ⦃a : T⦄, a ∈ l₁ → a ∈ l₂
|
|
|
|
|
|
|
|
|
|
infix `⊆`:50 := sublist
|
|
|
|
|
|
|
|
|
|
lemma nil_sub (l : list T) : [] ⊆ l :=
|
|
|
|
|
λ b i, false.elim (iff.mp (mem_nil b) i)
|
|
|
|
|
|
|
|
|
|
lemma sub.refl (l : list T) : l ⊆ l :=
|
|
|
|
|
λ b i, i
|
|
|
|
|
|
|
|
|
|
lemma sub.trans {l₁ l₂ l₃ : list T} (H₁ : l₁ ⊆ l₂) (H₂ : l₂ ⊆ l₃) : l₁ ⊆ l₃ :=
|
|
|
|
|
λ b i, H₂ (H₁ i)
|
|
|
|
|
|
|
|
|
|
lemma sub_cons (a : T) (l : list T) : l ⊆ a::l :=
|
|
|
|
|
λ b i, or.inr i
|
|
|
|
|
|
|
|
|
|
lemma cons_sub_cons {l₁ l₂ : list T} (a : T) (s : l₁ ⊆ l₂) : (a::l₁) ⊆ (a::l₂) :=
|
|
|
|
|
λ b Hin, or.elim Hin
|
|
|
|
|
(λ e : b = a, or.inl e)
|
|
|
|
|
(λ i : b ∈ l₁, or.inr (s i))
|
|
|
|
|
|
|
|
|
|
lemma sub_append_left (l₁ l₂ : list T) : l₁ ⊆ l₁++l₂ :=
|
|
|
|
|
λ b i, iff.mp' (mem_append_iff b l₁ l₂) (or.inl i)
|
|
|
|
|
|
|
|
|
|
lemma sub_append_right (l₁ l₂ : list T) : l₂ ⊆ l₁++l₂ :=
|
|
|
|
|
λ b i, iff.mp' (mem_append_iff b l₁ l₂) (or.inr i)
|
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- find -/
|
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
section
|
2014-10-12 20:06:00 +00:00
|
|
|
|
variable [H : decidable_eq T]
|
2014-10-05 20:09:56 +00:00
|
|
|
|
include H
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition find : T → list T → nat
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| a [] := 0
|
|
|
|
|
| a (b :: l) := if a = b then 0 else succ (find a l)
|
2014-09-08 04:06:32 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem find_nil (x : T) : find x [] = 0
|
2014-09-08 04:06:32 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem find_cons (x y : T) (l : list T) : find x (y::l) = if x = y then 0 else succ (find x l)
|
2014-09-08 04:06:32 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem find.not_mem {l : list T} {x : T} : ¬x ∈ l → find x l = length l :=
|
2015-02-11 20:49:27 +00:00
|
|
|
|
list.rec_on l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
(assume P₁ : ¬x ∈ [], _)
|
2014-09-08 04:06:32 +00:00
|
|
|
|
(take y l,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
assume iH : ¬x ∈ l → find x l = length l,
|
|
|
|
|
assume P₁ : ¬x ∈ y::l,
|
2015-03-29 20:59:36 +00:00
|
|
|
|
have P₂ : ¬(x = y ∨ x ∈ l), from iff.elim_right (not_iff_not_of_iff !mem_cons_iff) P₁,
|
2014-12-15 20:05:44 +00:00
|
|
|
|
have P₃ : ¬x = y ∧ ¬x ∈ l, from (iff.elim_left not_or_iff_not_and_not P₂),
|
2014-09-08 04:06:32 +00:00
|
|
|
|
calc
|
2015-01-07 21:38:11 +00:00
|
|
|
|
find x (y::l) = if x = y then 0 else succ (find x l) : !find_cons
|
2014-09-10 23:42:27 +00:00
|
|
|
|
... = succ (find x l) : if_neg (and.elim_left P₃)
|
|
|
|
|
... = succ (length l) : {iH (and.elim_right P₃)}
|
2015-01-07 21:38:11 +00:00
|
|
|
|
... = length (y::l) : !length_cons⁻¹)
|
2014-10-05 20:09:56 +00:00
|
|
|
|
end
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-12-23 22:34:16 +00:00
|
|
|
|
/- nth element -/
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2015-02-26 00:20:44 +00:00
|
|
|
|
definition nth [h : inhabited T] : list T → nat → T
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] n := arbitrary T
|
|
|
|
|
| (a :: l) 0 := a
|
|
|
|
|
| (a :: l) (n+1) := nth l n
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem nth_zero [h : inhabited T] (a : T) (l : list T) : nth (a :: l) 0 = a
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2015-01-07 21:38:11 +00:00
|
|
|
|
theorem nth_succ [h : inhabited T] (a : T) (l : list T) (n : nat) : nth (a::l) (n+1) = nth l n
|
2014-10-10 23:33:58 +00:00
|
|
|
|
|
2015-03-05 02:48:13 +00:00
|
|
|
|
open decidable
|
2015-03-05 04:30:19 +00:00
|
|
|
|
definition decidable_eq {A : Type} [H : decidable_eq A] : ∀ l₁ l₂ : list A, decidable (l₁ = l₂)
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] [] := inl rfl
|
|
|
|
|
| [] (b::l₂) := inr (λ H, list.no_confusion H)
|
|
|
|
|
| (a::l₁) [] := inr (λ H, list.no_confusion H)
|
|
|
|
|
| (a::l₁) (b::l₂) :=
|
2015-03-05 02:48:13 +00:00
|
|
|
|
match H a b with
|
|
|
|
|
| inl Hab :=
|
|
|
|
|
match decidable_eq l₁ l₂ with
|
|
|
|
|
| inl He := inl (eq.rec_on Hab (eq.rec_on He rfl))
|
|
|
|
|
| inr Hn := inr (λ H, list.no_confusion H (λ Hab Ht, absurd Ht Hn))
|
|
|
|
|
end
|
|
|
|
|
| inr Hnab := inr (λ H, list.no_confusion H (λ Hab Ht, absurd Hab Hnab))
|
|
|
|
|
end
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
section combinators
|
|
|
|
|
variables {A B C : Type}
|
|
|
|
|
|
|
|
|
|
definition map (f : A → B) : list A → list B
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := []
|
|
|
|
|
| (a :: l) := f a :: map l
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem map_nil (f : A → B) : map f [] = []
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
theorem map_cons (f : A → B) (a : A) (l : list A) : map f (a :: l) = f a :: map f l
|
|
|
|
|
|
|
|
|
|
theorem map_map (g : B → C) (f : A → B) : ∀ l : list A, map g (map f l) = map (g ∘ f) l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := rfl
|
|
|
|
|
| (a :: l) :=
|
2015-03-05 04:40:06 +00:00
|
|
|
|
show (g ∘ f) a :: map g (map f l) = map (g ∘ f) (a :: l),
|
|
|
|
|
by rewrite (map_map l)
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
theorem len_map (f : A → B) : ∀ l : list A, length (map f l) = length l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := rfl
|
|
|
|
|
| (a :: l) :=
|
2015-03-05 04:40:06 +00:00
|
|
|
|
show length (map f l) + 1 = length l + 1,
|
|
|
|
|
by rewrite (len_map l)
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
definition map₂ (f : A → B → C) : list A → list B → list C
|
|
|
|
|
| [] _ := []
|
|
|
|
|
| _ [] := []
|
|
|
|
|
| (x::xs) (y::ys) := f x y :: map₂ xs ys
|
|
|
|
|
|
2015-03-05 04:30:19 +00:00
|
|
|
|
definition foldl (f : A → B → A) : A → list B → A
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| a [] := a
|
|
|
|
|
| a (b :: l) := foldl (f a b) l
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
definition foldr (f : A → B → B) : B → list A → B
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| b [] := b
|
|
|
|
|
| b (a :: l) := f a (foldr b l)
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
section foldl_eq_foldr
|
|
|
|
|
-- foldl and foldr coincide when f is commutative and associative
|
|
|
|
|
parameters {α : Type} {f : α → α → α}
|
|
|
|
|
hypothesis (Hcomm : ∀ a b, f a b = f b a)
|
|
|
|
|
hypothesis (Hassoc : ∀ a b c, f a (f b c) = f (f a b) c)
|
|
|
|
|
include Hcomm Hassoc
|
|
|
|
|
|
|
|
|
|
theorem foldl_eq_of_comm_of_assoc : ∀ a b l, foldl f a (b::l) = f b (foldl f a l)
|
|
|
|
|
| a b nil := Hcomm a b
|
|
|
|
|
| a b (c::l) :=
|
|
|
|
|
begin
|
|
|
|
|
change (foldl f (f (f a b) c) l = f b (foldl f (f a c) l)),
|
|
|
|
|
rewrite -foldl_eq_of_comm_of_assoc,
|
|
|
|
|
change (foldl f (f (f a b) c) l = foldl f (f (f a c) b) l),
|
|
|
|
|
have H₁ : f (f a b) c = f (f a c) b, by rewrite [-Hassoc, -Hassoc, Hcomm b c],
|
|
|
|
|
rewrite H₁
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
theorem foldl_eq_foldr : ∀ a l, foldl f a l = foldr f a l
|
|
|
|
|
| a nil := rfl
|
|
|
|
|
| a (b :: l) :=
|
|
|
|
|
begin
|
|
|
|
|
rewrite foldl_eq_of_comm_of_assoc,
|
|
|
|
|
esimp,
|
|
|
|
|
change (f b (foldl f a l) = f b (foldr f a l)),
|
|
|
|
|
rewrite foldl_eq_foldr
|
|
|
|
|
end
|
|
|
|
|
end foldl_eq_foldr
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
definition all (p : A → Prop) (l : list A) : Prop :=
|
|
|
|
|
foldr (λ a r, p a ∧ r) true l
|
|
|
|
|
|
|
|
|
|
definition any (p : A → Prop) (l : list A) : Prop :=
|
|
|
|
|
foldr (λ a r, p a ∨ r) false l
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
definition decidable_all (p : A → Prop) [H : decidable_pred p] : ∀ l, decidable (all p l)
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := decidable_true
|
|
|
|
|
| (a :: l) :=
|
2015-03-05 04:30:19 +00:00
|
|
|
|
match H a with
|
|
|
|
|
| inl Hp₁ :=
|
|
|
|
|
match decidable_all l with
|
|
|
|
|
| inl Hp₂ := inl (and.intro Hp₁ Hp₂)
|
|
|
|
|
| inr Hn₂ := inr (not_and_of_not_right (p a) Hn₂)
|
|
|
|
|
end
|
|
|
|
|
| inr Hn := inr (not_and_of_not_left (all p l) Hn)
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
definition decidable_any (p : A → Prop) [H : decidable_pred p] : ∀ l, decidable (any p l)
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := decidable_false
|
|
|
|
|
| (a :: l) :=
|
2015-03-05 04:30:19 +00:00
|
|
|
|
match H a with
|
|
|
|
|
| inl Hp := inl (or.inl Hp)
|
|
|
|
|
| inr Hn₁ :=
|
|
|
|
|
match decidable_any l with
|
|
|
|
|
| inl Hp₂ := inl (or.inr Hp₂)
|
|
|
|
|
| inr Hn₂ := inr (not_or Hn₁ Hn₂)
|
|
|
|
|
end
|
|
|
|
|
end
|
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
definition zip (l₁ : list A) (l₂ : list B) : list (A × B) :=
|
|
|
|
|
map₂ (λ a b, (a, b)) l₁ l₂
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
definition unzip : list (A × B) → list A × list B
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := ([], [])
|
|
|
|
|
| ((a, b) :: l) :=
|
2015-03-05 04:30:19 +00:00
|
|
|
|
match unzip l with
|
|
|
|
|
| (la, lb) := (a :: la, b :: lb)
|
|
|
|
|
end
|
|
|
|
|
|
2015-03-14 05:25:21 +00:00
|
|
|
|
theorem unzip_nil : unzip (@nil (A × B)) = ([], [])
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
2015-03-05 04:40:06 +00:00
|
|
|
|
theorem unzip_cons (a : A) (b : B) (l : list (A × B)) :
|
|
|
|
|
unzip ((a, b) :: l) = match unzip l with (la, lb) := (a :: la, b :: lb) end
|
2015-03-05 04:30:19 +00:00
|
|
|
|
|
|
|
|
|
theorem zip_unzip : ∀ (l : list (A × B)), zip (pr₁ (unzip l)) (pr₂ (unzip l)) = l
|
2015-03-14 05:25:21 +00:00
|
|
|
|
| [] := rfl
|
|
|
|
|
| ((a, b) :: l) :=
|
2015-03-05 04:30:19 +00:00
|
|
|
|
begin
|
|
|
|
|
rewrite unzip_cons,
|
|
|
|
|
have r : zip (pr₁ (unzip l)) (pr₂ (unzip l)) = l, from zip_unzip l,
|
|
|
|
|
revert r,
|
|
|
|
|
apply (prod.cases_on (unzip l)),
|
2015-03-28 00:26:06 +00:00
|
|
|
|
intros [la, lb, r],
|
2015-03-05 04:30:19 +00:00
|
|
|
|
rewrite -r
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
end combinators
|
|
|
|
|
|
2015-03-14 20:06:09 +00:00
|
|
|
|
/- flat -/
|
|
|
|
|
variable {A : Type}
|
|
|
|
|
|
|
|
|
|
definition flat (l : list (list A)) : list A :=
|
|
|
|
|
foldl append nil l
|
|
|
|
|
|
2014-08-05 00:07:59 +00:00
|
|
|
|
end list
|
2015-03-05 02:48:13 +00:00
|
|
|
|
|
2015-03-05 04:40:06 +00:00
|
|
|
|
attribute list.decidable_eq [instance]
|
2015-03-05 02:48:13 +00:00
|
|
|
|
attribute list.decidable_mem [instance]
|
2015-03-05 04:30:19 +00:00
|
|
|
|
attribute list.decidable_any [instance]
|
|
|
|
|
attribute list.decidable_all [instance]
|