2014-07-30 00:04:25 +00:00
|
|
|
|
----------------------------------------------------------------------------------------------------
|
|
|
|
|
--- Copyright (c) 2014 Parikshit Khanna. All rights reserved.
|
|
|
|
|
--- Released under Apache 2.0 license as described in the file LICENSE.
|
2014-10-05 20:09:56 +00:00
|
|
|
|
--- Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura
|
2014-07-30 00:04:25 +00:00
|
|
|
|
----------------------------------------------------------------------------------------------------
|
2014-10-20 22:32:42 +00:00
|
|
|
|
import logic tools.helper_tactics tools.tactic data.nat.basic
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- Theory list
|
|
|
|
|
-- ===========
|
|
|
|
|
--
|
|
|
|
|
-- Basic properties of lists.
|
2014-10-05 20:09:56 +00:00
|
|
|
|
open eq.ops helper_tactics nat
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
inductive list (T : Type) : Type :=
|
2014-08-22 22:46: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
|
2014-10-19 18:13:36 +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
|
|
|
|
|
|
|
|
|
-- Concat
|
|
|
|
|
-- ------
|
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition append (s t : list T) : list T :=
|
2014-09-10 23:42:27 +00:00
|
|
|
|
rec t (λx l u, x::u) s
|
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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem append.nil_left (t : list T) : nil ++ t = t
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem append.nil_right (t : list T) : t ++ nil = t :=
|
2014-09-10 23:42:27 +00:00
|
|
|
|
induction_on t rfl (λx l H, H ▸ rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem append.assoc (s t u : list T) : s ++ t ++ u = s ++ (t ++ u) :=
|
2014-09-10 23:42:27 +00:00
|
|
|
|
induction_on s rfl (λx l H, H ▸ rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- Length
|
|
|
|
|
-- ------
|
|
|
|
|
|
2014-09-10 23:42:27 +00:00
|
|
|
|
definition length : list T → nat :=
|
2014-09-10 14:55:32 +00:00
|
|
|
|
rec 0 (λx l m, succ m)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem length.nil : length (@nil T) = 0
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem length.cons (x : T) (t : list T) : length (x::t) = succ (length t)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem length.append (s t : list T) : length (s ++ t) = length s + length t :=
|
2014-10-02 01:39:47 +00:00
|
|
|
|
induction_on s (!add.zero_left⁻¹) (λx s H, !add.succ_left⁻¹ ▸ H ▸ rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- add_rewrite length_nil length_cons
|
|
|
|
|
|
|
|
|
|
-- Append
|
|
|
|
|
-- ------
|
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition concat (x : T) : list T → list T :=
|
2014-09-10 23:42:27 +00:00
|
|
|
|
rec [x] (λy l l', y::l')
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem concat.nil (x : T) : concat x nil = [x]
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem concat.eq_append (x : T) (l : list T) : concat x l = l ++ [x]
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
|
|
|
|
-- add_rewrite append_nil append_cons
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- Reverse
|
|
|
|
|
-- -------
|
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition reverse : list T → list T :=
|
|
|
|
|
rec nil (λx l r, r ++ [x])
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem reverse.nil : reverse (@nil T) = nil
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem reverse.cons (x : T) (l : list T) : reverse (x::l) = concat x (reverse l)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem reverse.singleton (x : T) : reverse [x] = [x]
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem reverse.append (s t : list T) : reverse (s ++ t) = (reverse t) ++ (reverse s) :=
|
|
|
|
|
induction_on s (!append.nil_right⁻¹)
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(λx s H, calc
|
|
|
|
|
reverse (x::s ++ t) = reverse t ++ reverse s ++ [x] : {H}
|
2014-10-05 20:09:56 +00:00
|
|
|
|
... = reverse t ++ (reverse s ++ [x]) : !append.assoc)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem reverse.reverse (l : list T) : reverse (reverse l) = l :=
|
|
|
|
|
induction_on l rfl (λx l' H, H ▸ !reverse.append)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem concat.eq_reverse_cons (x : T) (l : list T) : concat x l = reverse (x :: reverse l) :=
|
2014-09-10 23:42:27 +00:00
|
|
|
|
induction_on l rfl
|
|
|
|
|
(λy l' H, calc
|
2014-10-05 20:09:56 +00:00
|
|
|
|
concat x (y::l') = (y::l') ++ [x] : !concat.eq_append
|
|
|
|
|
... = reverse (reverse (y::l')) ++ [x] : {!reverse.reverse⁻¹})
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- Head and tail
|
|
|
|
|
-- -------------
|
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition head (x : T) : list T → T :=
|
|
|
|
|
rec x (λx l h, x)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem head.nil (x : T) : head x nil = x
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem head.cons (x x' : T) (t : list T) : head x' (x::t) = x
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem head.concat {s : list T} (t : list T) (x : T) : s ≠ nil → (head x (s ++ t) = head x s) :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
cases_on s
|
2014-09-02 02:44:04 +00:00
|
|
|
|
(take H : nil ≠ nil, absurd rfl H)
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(take x s, take H : x::s ≠ nil,
|
2014-07-30 00:04:25 +00:00
|
|
|
|
calc
|
2014-10-05 20:09:56 +00:00
|
|
|
|
head x (x::s ++ t) = head x (x::(s ++ t)) : {!append.cons}
|
|
|
|
|
... = x : !head.cons
|
|
|
|
|
... = head x (x::s) : !head.cons⁻¹)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition tail : list T → list T :=
|
|
|
|
|
rec nil (λx l b, l)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem tail.nil : tail (@nil T) = nil
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem tail.cons (x : T) (l : list T) : tail (x::l) = l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem cons_head_tail {l : list T} (x : T) : l ≠ nil → (head x l)::(tail l) = l :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
cases_on l
|
2014-09-02 02:44:04 +00:00
|
|
|
|
(assume H : nil ≠ nil, absurd rfl H)
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(take x l, assume H : x::l ≠ nil, rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- List membership
|
|
|
|
|
-- ---------------
|
|
|
|
|
|
2014-09-10 14:55:32 +00:00
|
|
|
|
definition mem (x : T) : list T → Prop :=
|
|
|
|
|
rec false (λy l H, x = y ∨ H)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-21 21:08:07 +00:00
|
|
|
|
notation e ∈ s := mem e s
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.nil (x : T) : x ∈ nil ↔ false :=
|
2014-09-10 14:55:32 +00:00
|
|
|
|
iff.rfl
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.cons (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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.concat_imp_or {x : T} {s t : list T} : x ∈ s ++ t → x ∈ s ∨ x ∈ t :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
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-09-05 04:25:21 +00:00
|
|
|
|
have H3 : x = y ∨ x ∈ s ∨ x ∈ t, from or.imp_or_right H2 IH,
|
|
|
|
|
iff.elim_right or.assoc H3)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.or_imp_concat {x : T} {s t : list T} : x ∈ s ∨ x ∈ t → x ∈ s ++ t :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on s
|
2014-09-05 04:25:21 +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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.concat (x : T) (s t : list T) : x ∈ s ++ t ↔ x ∈ s ∨ x ∈ t :=
|
|
|
|
|
iff.intro mem.concat_imp_or mem.or_imp_concat
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem mem.split {x : T} {l : list T} : x ∈ l → ∃s t : list T, l = s ++ (x::t) :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on l
|
2014-10-05 20:09:56 +00:00
|
|
|
|
(take H : x ∈ nil, 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,
|
2014-10-05 20:09:56 +00:00
|
|
|
|
exists_intro nil (!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-10-05 20:09:56 +00:00
|
|
|
|
!exists_intro (!exists_intro H4)))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-10-12 20:06:00 +00:00
|
|
|
|
definition mem.is_decidable [instance] (H : decidable_eq T) (x : T) (l : list T) : decidable (x ∈ l) :=
|
2014-09-08 04:06:32 +00:00
|
|
|
|
rec_on l
|
2014-10-05 20:09:56 +00:00
|
|
|
|
(decidable.inr (iff.false_elim !mem.nil))
|
|
|
|
|
(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
|
2014-10-05 20:09:56 +00:00
|
|
|
|
iff.elim_right (iff.flip_sign !mem.cons) H1,
|
2014-09-08 04:06:32 +00:00
|
|
|
|
decidable.inr H2)))
|
|
|
|
|
|
2014-07-31 20:33:35 +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
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
definition find (x : T) : list T → nat :=
|
2014-09-10 14:55:32 +00:00
|
|
|
|
rec 0 (λy l b, if x = y then 0 else succ b)
|
2014-09-08 04:06:32 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem find.nil (x : T) : find x nil = 0
|
2014-09-08 04:06:32 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +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 :=
|
2014-09-08 04:06:32 +00:00
|
|
|
|
rec_on l
|
2014-09-10 23:42:27 +00:00
|
|
|
|
(assume P₁ : ¬x ∈ nil, rfl)
|
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,
|
2014-10-05 20:09:56 +00:00
|
|
|
|
have P₂ : ¬(x = y ∨ x ∈ l), from iff.elim_right (iff.flip_sign !mem.cons) P₁,
|
2014-09-10 23:42:27 +00:00
|
|
|
|
have P₃ : ¬x = y ∧ ¬x ∈ l, from (iff.elim_left not_or P₂),
|
2014-09-08 04:06:32 +00:00
|
|
|
|
calc
|
2014-10-05 20:09:56 +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₃)}
|
2014-10-05 20:09:56 +00:00
|
|
|
|
... = length (y::l) : !length.cons⁻¹)
|
|
|
|
|
end
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- nth element
|
|
|
|
|
-- -----------
|
|
|
|
|
|
2014-09-10 23:42:27 +00:00
|
|
|
|
definition nth (x : T) (l : list T) (n : nat) : T :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
nat.rec (λl, head x l) (λm f l, f (tail l)) n l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem nth.zero (x : T) (l : list T) : nth x l 0 = head x l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-10-05 20:09:56 +00:00
|
|
|
|
theorem nth.succ (x : T) (l : list T) (n : nat) : nth x l (succ n) = nth x (tail l) n
|
2014-10-10 23:33:58 +00:00
|
|
|
|
|
2014-08-05 00:07:59 +00:00
|
|
|
|
end list
|