lean2/tests/lean/run/tree_height.lean
Leonardo de Moura 064ecd3e3d refactor(library/data/nat): declare lt and le asap using inductive definitions, and make key theorems transparent for definitional package
We also define key theorems that will be used to generate the
automatically generated a well-founded subterm relation for inductive
datatypes.
We also prove decidability and wf theorems asap.
2014-11-22 00:19:39 -08:00

36 lines
1.1 KiB
Text

import logic data.nat
open eq.ops nat
inductive tree (A : Type) :=
leaf : A → tree A,
node : tree A → tree A → tree A
namespace tree
definition height {A : Type} (t : tree A) : nat :=
rec_on t
(λ a, zero)
(λ t₁ t₂ h₁ h₂, succ (max h₁ h₂))
definition height_lt {A : Type} : tree A → tree A → Prop :=
inv_image lt (@height A)
definition height_lt.wf (A : Type) : well_founded (@height_lt A) :=
inv_image.wf height lt.wf
theorem height_lt.node_left {A : Type} (t₁ t₂ : tree A) : height_lt t₁ (node t₁ t₂) :=
le_imp_lt_succ (max.left (height t₁) (height t₂))
theorem height_lt.node_right {A : Type} (t₁ t₂ : tree A) : height_lt t₂ (node t₁ t₂) :=
le_imp_lt_succ (max.right (height t₁) (height t₂))
theorem height_lt.trans {A : Type} : transitive (@height_lt A) :=
inv_image.trans lt height @lt.trans
example : height_lt (leaf 2) (node (leaf 1) (leaf 2)) :=
!height_lt.node_right
example : height_lt (leaf 2) (node (node (leaf 1) (leaf 2)) (leaf 3)) :=
height_lt.trans !height_lt.node_right !height_lt.node_left
end tree