File: bug_1776.v

package info (click to toggle)
coq 9.1.0%2Bdfsg-3
  • links: PTS, VCS
  • area: main
  • in suites: experimental
  • size: 35,964 kB
  • sloc: ml: 239,908; sh: 4,355; python: 2,985; ansic: 2,644; makefile: 874; lisp: 171; javascript: 63; xml: 24; sed: 2
file content (22 lines) | stat: -rw-r--r-- 620 bytes parent folder | download | duplicates (13)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
Axiom pair : nat -> nat -> nat -> Prop.
Axiom pl : (nat -> Prop) -> (nat -> Prop) -> (nat -> Prop).
Axiom plImpR : forall k P Q,
  pl P Q k -> forall (Q':nat -> Prop),
    (forall k', Q k' -> Q' k') ->
    pl P Q' k.

Definition nexists (P:nat -> nat -> Prop) : nat -> Prop :=
  fun k' => exists k, P k k'.

Goal forall a A m,
  True ->
  (pl A (nexists (fun x => (nexists
    (fun y => pl (pair a (S x)) (pair a (S y))))))) m.
Proof.
  intros.
  eapply plImpR; [ | intros; econstructor; econstructor; eauto].
  clear H;
    match goal with
       | |- (pl _ (pl (pair _ ?x) _)) _ => replace x with 0
    end.
Admitted.