File: InductiveVsImplicitsVsTC.v

package info (click to toggle)
coq 8.20.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: forky, sid, trixie
  • size: 44,116 kB
  • sloc: ml: 234,160; sh: 4,301; python: 3,270; ansic: 2,644; makefile: 882; lisp: 172; javascript: 63; xml: 24; sed: 2
file content (26 lines) | stat: -rw-r--r-- 700 bytes parent folder | download | duplicates (2)
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
Module NoConv.
  Class C := {}.

  Definition useC {c:C} := nat.

  Inductive foo {a b : C} := CC : useC -> foo.
  (* If TC search runs before parameter unification it will pick the
     wrong instance for the first parameter.

     useC makes sure we don't completely skip TC search.
   *)
End NoConv.

Module ForConv.

  Class Bla := { bla : Type }.

  #[export] Instance bli : Bla := { bla := nat }.

  Inductive vs := C : forall x : bla, x = 2 -> vs.
  (* here we need to resolve TC to pass the conversion problem if we
     combined with the previous example it would fail as TC resolution
     for conversion is unrestricted and so would resolve the
     conclusion too early. *)

End ForConv.