File: nested.v

package info (click to toggle)
coq-doc 8.16.1-1
  • links: PTS, VCS
  • area: non-free
  • in suites: bookworm
  • size: 42,788 kB
  • sloc: ml: 219,673; sh: 4,035; python: 3,372; ansic: 2,529; makefile: 728; lisp: 279; javascript: 87; xml: 24; sed: 2
file content (47 lines) | stat: -rw-r--r-- 813 bytes parent folder | download | duplicates (4)
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
From Coq Require Import Uint63 PArray.

Open Scope array_scope.

Module OneLevel.

Inductive foo : Set :=
  C : array foo -> foo.

Fixpoint f1 (x : foo) {struct x} : False :=
  match x with
  | C t => f1 (t.[0])
  end.

Fixpoint f2 (x : foo) {struct x} : False :=
  f2 match x with
  | C t => t.[0]
  end.

Fixpoint f3 (x : foo) {struct x} : False :=
  match x with
  | C t => f3 (PArray.default t)
  end.

End OneLevel.

Module TwoLevels.

Inductive foo : Set :=
  C : array (array foo) -> foo.

Fixpoint f1 (x : foo) {struct x} : False :=
  match x with
  | C t => f1 (t.[0].[0])
  end.

Fixpoint f2 (x : foo) {struct x} : False :=
  f2 match x with
  | C t => t.[0].[0]
  end.

Fixpoint f3 (x : foo) {struct x} : False :=
  match x with
  | C t => f3 (PArray.default (PArray.default t))
  end.

End TwoLevels.