1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71
|
Inductive foo : nat -> Type :=
| fz : foo 0
| fs : forall {n}, foo (S n).
Definition test (n : nat) : foo n -> nat.
refine (
match n with
| 0 =>
fun f =>
match f in foo n return match n with 0 => nat | S _ => unit end with
| fz => 0
| fs => tt
end
| S n =>
fun f =>
match f in foo n return match n with 0 => unit | S _ => nat end with
| fz => tt
| fs => 42
end
end).
Qed.
Definition test2 n : foo n -> nat.
refine (
match n with
| 0 =>
fun (f: foo 0) =>
match f in foo n return match n with 0 => nat | S _ => unit end with
| fz => 0
| fs => tt
end
| S n =>
fun f =>
match f in foo n return match n with 0 => unit | S _ => nat end with
| fz => tt
| fs => 42
end
end).
Qed.
Definition test3 (n : nat) : foo n -> nat :=
match n with
| 0 =>
fun f =>
match f in foo n return match n with 0 => nat | S _ => unit end with
| fz => 0
| fs => tt
end
| S n =>
fun f =>
match f in foo n return match n with 0 => unit | S _ => nat end with
| fz => tt
| fs => 42
end
end.
Definition test4 (n : nat) : bool -> foo n -> nat :=
match n with
| 0 =>
fun b f =>
match f in foo n return match n with 0 => nat | S _ => unit end with
| fz => 0
| fs => tt
end
| S _ =>
fun b f =>
match f in foo n return match n with 0 => unit | S _ => nat end with
| fz => tt
| fs => 42
end
end.
|