File: bug_4709.v

package info (click to toggle)
coq-doc 8.20.0-2
  • links: PTS, VCS
  • area: non-free
  • in suites: forky, sid, trixie
  • size: 46,708 kB
  • sloc: ml: 234,429; sh: 4,686; python: 3,359; ansic: 2,644; makefile: 842; lisp: 172; javascript: 87; xml: 24; sed: 2
file content (18 lines) | stat: -rw-r--r-- 481 bytes parent folder | download | duplicates (6)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18

(** Bug 4709 https://coq.inria.fr/bug/4709
    Extraction wasn't reducing primitive projections in types. *)

Require Extraction.

Set Primitive Projections.

Record t := Foo { foo : Type }.
Definition ty := foo (Foo nat).

(* Without proper reduction of primitive projections in
   [extract_type], the type [ty] was extracted as [Tunknown].
   Let's check it isn't the case anymore. *)

Parameter check : nat.
Extract Constant check => "(O:ty)".
Extraction TestCompile ty check.