lean2/tests/lean/run/coe10.lean
Leonardo de Moura edcbe6fe10 feat(frontends/lean): allow multiple coercions from class A to B, closes #187
See new tests (for examples)
tests/lean/run/coe10.lean
tests/lean/run/coe11.lean
tests/lean/run/coe9.lean

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-09-14 12:59:43 -07:00

17 lines
377 B
Text

import data.nat
open nat
inductive fn2 (A B C : Type) :=
mk : (A → C) → (B → C) → fn2 A B C
definition to_ac [coercion] {A B C : Type} (f : fn2 A B C) : A → C :=
fn2.rec (λ f g, f) f
definition to_bc [coercion] {A B C : Type} (f : fn2 A B C) : B → C :=
fn2.rec (λ f g, g) f
variable f : fn2 Prop nat nat
variable a : Prop
variable n : nat
check f a
check f n