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.
|
|
|
|
|
--- Authors: Parikshit Khanna, Jeremy Avigad
|
|
|
|
|
----------------------------------------------------------------------------------------------------
|
|
|
|
|
|
|
|
|
|
-- Theory list
|
|
|
|
|
-- ===========
|
|
|
|
|
--
|
|
|
|
|
-- Basic properties of lists.
|
|
|
|
|
|
2014-08-01 00:48:51 +00:00
|
|
|
|
import tools.tactic
|
|
|
|
|
import data.nat
|
2014-09-02 02:44:04 +00:00
|
|
|
|
import logic tools.helper_tactics
|
2014-08-01 00:48:51 +00:00
|
|
|
|
-- import if -- for find
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open nat
|
|
|
|
|
open eq_ops
|
|
|
|
|
open helper_tactics
|
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
|
|
|
|
|
|
|
|
|
|
-- Type
|
|
|
|
|
-- ----
|
2014-08-22 23:36:47 +00:00
|
|
|
|
infix `::` := cons
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
|
|
|
|
|
variable {T : Type}
|
|
|
|
|
|
2014-09-03 22:13:03 +00:00
|
|
|
|
theorem induction_on [protected] {P : list T → Prop} (l : list T) (Hnil : P nil)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(Hind : forall x : T, forall l : list T, forall H : P l, P (cons x l)) : P l :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec Hnil Hind l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-03 22:13:03 +00:00
|
|
|
|
theorem cases_on [protected] {P : list T → Prop} (l : list T) (Hnil : P nil)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(Hcons : forall x : T, forall l : list T, P (cons x l)) : P l :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on l Hnil (take x l IH, Hcons x l)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
notation `[` l:(foldr `,` (h t, cons h t) nil) `]` := l
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- Concat
|
|
|
|
|
-- ------
|
|
|
|
|
|
|
|
|
|
definition concat (s t : list T) : list T :=
|
2014-09-04 22:03:59 +00:00
|
|
|
|
rec t (fun x : T, fun l : list T, fun u : list T, cons x u) s
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
infixl `++` : 65 := concat
|
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem nil_concat {t : list T} : nil ++ t = t
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem cons_concat {x : T} {s t : list T} : (x :: s) ++ t = x :: (s ++ t)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem concat_nil {t : list T} : t ++ nil = t :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on t rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(take (x : T) (l : list T) (H : concat l nil = l),
|
2014-09-02 02:44:04 +00:00
|
|
|
|
show concat (cons x l) nil = cons x l, from H ▸ rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem concat_assoc {s t u : list T} : s ++ t ++ u = s ++ (t ++ u) :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on s rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(take x l,
|
|
|
|
|
assume H : concat (concat l t) u = concat l (concat t u),
|
|
|
|
|
calc
|
2014-09-02 02:44:04 +00:00
|
|
|
|
concat (concat (cons x l) t) u = cons x (concat (concat l t) u) : rfl
|
|
|
|
|
... = cons x (concat l (concat t u)) : {H}
|
|
|
|
|
... = concat (cons x l) (concat t u) : rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- Length
|
|
|
|
|
-- ------
|
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition length : list T → ℕ := rec 0 (fun x l m, succ m)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-08-27 01:47:36 +00:00
|
|
|
|
theorem length_nil : length (@nil T) = 0 := rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +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-09-02 02:44:04 +00:00
|
|
|
|
theorem length_concat {s t : list T} : length (s ++ t) = length s + length t :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on s
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(calc
|
2014-08-27 01:47:36 +00:00
|
|
|
|
length (concat nil t) = length t : rfl
|
|
|
|
|
... = zero + length t : {add_zero_left⁻¹}
|
|
|
|
|
... = length (@nil T) + length t : rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(take x s,
|
|
|
|
|
assume H : length (concat s t) = length s + length t,
|
|
|
|
|
calc
|
2014-08-27 01:47:36 +00:00
|
|
|
|
length (concat (cons x s) t ) = succ (length (concat s t)) : rfl
|
2014-09-02 02:44:04 +00:00
|
|
|
|
... = succ (length s + length t) : {H}
|
2014-08-27 01:47:36 +00:00
|
|
|
|
... = succ (length s) + length t : {add_succ_left⁻¹}
|
|
|
|
|
... = length (cons x 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
|
|
|
|
|
|
|
|
|
|
-- Append
|
|
|
|
|
-- ------
|
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition append (x : T) : list T → list T := rec [x] (fun y l l', y :: l')
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem append_nil {x : T} : append x nil = [x]
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem append_cons {x y : T} {l : list T} : append x (y :: l) = y :: (append x l)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem append_eq_concat {x : T} {l : list T} : append 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-04 22:03:59 +00:00
|
|
|
|
definition reverse : list T → list T := rec nil (fun x l r, r ++ [x])
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem reverse_nil : reverse (@nil T) = nil
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem reverse_cons {x : T} {l : list T} : reverse (x :: l) = append x (reverse l)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem reverse_singleton {x : T} : reverse [x] = [x]
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem reverse_concat {s t : list T} : reverse (s ++ t) = (reverse t) ++ (reverse s) :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on s (concat_nil⁻¹)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(take x s,
|
|
|
|
|
assume IH : reverse (s ++ t) = concat (reverse t) (reverse s),
|
2014-07-30 00:04:25 +00:00
|
|
|
|
calc
|
2014-09-02 02:44:04 +00:00
|
|
|
|
reverse ((x :: s) ++ t) = reverse (s ++ t) ++ [x] : rfl
|
|
|
|
|
... = reverse t ++ reverse s ++ [x] : {IH}
|
|
|
|
|
... = reverse t ++ (reverse s ++ [x]) : concat_assoc
|
|
|
|
|
... = reverse t ++ (reverse (x :: s)) : rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem reverse_reverse {l : list T} : reverse (reverse l) = l :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on l rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(take x l',
|
|
|
|
|
assume H: reverse (reverse l') = l',
|
2014-07-31 20:33:35 +00:00
|
|
|
|
show reverse (reverse (x :: l')) = x :: l', from
|
2014-07-30 00:04:25 +00:00
|
|
|
|
calc
|
2014-09-02 02:44:04 +00:00
|
|
|
|
reverse (reverse (x :: l')) = reverse (reverse l' ++ [x]) : rfl
|
|
|
|
|
... = reverse [x] ++ reverse (reverse l') : reverse_concat
|
|
|
|
|
... = [x] ++ l' : {H}
|
|
|
|
|
... = x :: l' : rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem append_eq_reverse_cons {x : T} {l : list T} : append x l = reverse (x :: reverse l) :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
induction_on l rfl
|
2014-07-30 00:04:25 +00:00
|
|
|
|
(take y l',
|
|
|
|
|
assume H : append x l' = reverse (x :: reverse l'),
|
|
|
|
|
calc
|
2014-09-02 02:44:04 +00:00
|
|
|
|
append x (y :: l') = (y :: l') ++ [ x ] : append_eq_concat
|
|
|
|
|
... = concat (reverse (reverse (y :: l'))) [ x ] : {reverse_reverse⁻¹}
|
|
|
|
|
... = reverse (x :: (reverse (y :: l'))) : rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- Head and tail
|
|
|
|
|
-- -------------
|
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition head (x : T) : list T → T := rec x (fun x l h, x)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem head_nil {x : T} : head x (@nil T) = x
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +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-09-02 02:44:04 +00:00
|
|
|
|
theorem head_concat {s 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)
|
|
|
|
|
(take x s, take H : cons x s ≠ nil,
|
2014-07-30 00:04:25 +00:00
|
|
|
|
calc
|
2014-09-02 02:44:04 +00:00
|
|
|
|
head x (concat (cons x s) t) = head x (cons x (concat s t)) : {cons_concat}
|
|
|
|
|
... = x : {head_cons}
|
|
|
|
|
... = head x (cons x s) : {head_cons⁻¹})
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition tail : list T → list T := rec nil (fun x l b, l)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem tail_nil : tail (@nil T) = nil
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem tail_cons {x : T} {l : list T} : tail (cons x l) = l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem cons_head_tail {x : T} {l : list 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)
|
|
|
|
|
(take x l, assume H : cons x l ≠ nil, rfl)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- List membership
|
|
|
|
|
-- ---------------
|
|
|
|
|
|
2014-09-04 22:03:59 +00:00
|
|
|
|
definition mem (x : T) : list T → Prop := rec false (fun y l H, x = y ∨ H)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-08-22 23:36:47 +00:00
|
|
|
|
infix `∈` := mem
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- TODO: constructively, equality is stronger. Use that?
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem mem_nil {x : T} : x ∈ nil ↔ false := iff_rfl
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem mem_cons {x y : T} {l : list T} : mem x (cons y l) ↔ (x = y ∨ mem x l) := iff_rfl
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
theorem mem_concat_imp_or {x : T} {s t : list T} : x ∈ s ++ t → x ∈ s ∨ x ∈ t :=
|
2014-09-03 22:13:03 +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,
|
|
|
|
|
assume H1 : x ∈ (y :: s) ++ t,
|
|
|
|
|
have H2 : x = y ∨ x ∈ s ++ t, from H1,
|
2014-08-22 23:36:47 +00:00
|
|
|
|
have H3 : x = y ∨ x ∈ s ∨ x ∈ t, from or_imp_or_right H2 IH,
|
2014-08-28 18:10:04 +00:00
|
|
|
|
iff_elim_right or_assoc H3)
|
2014-07-31 20:33:35 +00:00
|
|
|
|
|
2014-09-02 02:44:04 +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-02 02:44:04 +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,
|
|
|
|
|
assume H : x ∈ y :: s ∨ x ∈ t,
|
|
|
|
|
or_elim H
|
|
|
|
|
(assume H1,
|
|
|
|
|
or_elim H1
|
2014-08-28 18:10:04 +00:00
|
|
|
|
(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-09-02 02:44:04 +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-09-02 02:44:04 +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-09-02 02:44:04 +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,
|
|
|
|
|
assume IH : x ∈ l → ∃s t : list T, l = s ++ (x :: t),
|
|
|
|
|
assume H : x ∈ y :: l,
|
|
|
|
|
or_elim H
|
|
|
|
|
(assume H1 : x = y,
|
|
|
|
|
exists_intro nil
|
2014-09-02 02:44:04 +00:00
|
|
|
|
(exists_intro l (subst H1 rfl)))
|
2014-07-31 20:33:35 +00:00
|
|
|
|
(assume H1 : x ∈ l,
|
|
|
|
|
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-02 02:44:04 +00:00
|
|
|
|
from subst H3 rfl,
|
2014-07-31 20:33:35 +00:00
|
|
|
|
exists_intro _ (exists_intro _ H4)))
|
|
|
|
|
|
|
|
|
|
-- Find
|
|
|
|
|
-- ----
|
|
|
|
|
|
|
|
|
|
-- to do this: need decidability of = for nat
|
|
|
|
|
-- definition find (x : T) : list T → nat
|
2014-09-04 22:03:59 +00:00
|
|
|
|
-- := rec 0 (fun y l b, if x = y then 0 else succ b)
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
|
|
|
|
-- theorem find_nil (f : T) : find f nil = 0
|
|
|
|
|
-- :=refl _
|
|
|
|
|
|
|
|
|
|
-- theorem find_cons (x y : T) (l : list T) : find x (cons y l) =
|
|
|
|
|
-- if x = y then 0 else succ (find x l)
|
|
|
|
|
-- := refl _
|
|
|
|
|
|
|
|
|
|
-- theorem not_mem_find (l : list T) (x : T) : ¬ mem x l → find x l = length l
|
|
|
|
|
-- :=
|
2014-09-03 22:13:03 +00:00
|
|
|
|
-- @induction_on T (λl, ¬ mem x l → find x l = length l) l
|
|
|
|
|
-- -- induction_on l
|
2014-07-30 00:04:25 +00:00
|
|
|
|
-- (assume P1 : ¬ mem x nil,
|
|
|
|
|
-- show find x nil = length nil, from
|
|
|
|
|
-- calc
|
|
|
|
|
-- find x nil = 0 : find_nil _
|
|
|
|
|
-- ... = length nil : by simp)
|
|
|
|
|
-- (take y l,
|
|
|
|
|
-- assume IH : ¬ (mem x l) → find x l = length l,
|
|
|
|
|
-- assume P1 : ¬ (mem x (cons y l)),
|
|
|
|
|
-- have P2 : ¬ (mem x l ∨ (y = x)), from subst P1 (mem_cons _ _ _),
|
|
|
|
|
-- have P3 : ¬ (mem x l) ∧ (y ≠ x),from subst P2 (not_or _ _),
|
|
|
|
|
-- have P4 : x ≠ y, from ne_symm (and_elim_right P3),
|
|
|
|
|
-- calc
|
|
|
|
|
-- find x (cons y l) = succ (find x l) :
|
|
|
|
|
-- trans (find_cons _ _ _) (not_imp_if_eq P4 _ _)
|
|
|
|
|
-- ... = succ (length l) : {IH (and_elim_left P3)}
|
|
|
|
|
-- ... = length (cons y l) : symm (length_cons _ _))
|
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- nth element
|
|
|
|
|
-- -----------
|
|
|
|
|
|
2014-09-02 02:44:04 +00:00
|
|
|
|
definition nth (x : T) (l : list T) (n : ℕ) : 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-09-02 02:44:04 +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-09-02 02:44:04 +00:00
|
|
|
|
theorem nth_succ {x : T} {l : list T} {n : ℕ} : nth x l (succ n) = nth x (tail l) n
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
end
|
2014-07-30 00:04:25 +00:00
|
|
|
|
|
2014-07-31 20:33:35 +00:00
|
|
|
|
-- declare global notation outside the section
|
2014-08-22 23:36:47 +00:00
|
|
|
|
infixl `++` := concat
|
2014-08-08 00:52:51 +00:00
|
|
|
|
|
2014-08-05 00:07:59 +00:00
|
|
|
|
end list
|