Equality between functional and inductive definitions

拟墨画扇 提交于 2020-01-04 04:15:08

问题


I have an inductive definition of the proposition P (or repeats l) that a lists contains repeating elements, and a functional definition of it's negation Q (or no_repeats l).

I want to show that P <-> ~ Q and ~ P <-> Q. I have been able to show three of the four implications, but ~ Q -> P seems to be different, because I'm unable to extract data from ~Q.

Require Import List.
Variable A : Type.

Inductive repeats : list A -> Prop :=            (*   repeats  *)
  repeats_hd l x : In x l    -> repeats (x::l)
| repeats_tl l x : repeats l -> repeats (x::l).

Fixpoint no_repeats (l: list A): Prop :=
  match l with nil => True | a::l' =>  ~ In a l' /\ no_repeats l' end.

Lemma not_no_repeats_repeats: forall l, (~ no_repeats l) -> repeats l.
  induction l; simpl. tauto. intros.

After doing induction on l, the second case is

  IHl : ~ no_repeats l -> repeats l
  H : ~ (~ In a l /\ no_repeats l)
  ============================
  repeats (a :: l)

Is it possible to deduce In a l \/ ~ no_repeats l (which is sufficient) from this?


回答1:


Your statement implies that equality on A supports double negation elimination:

Require Import List.
Import ListNotations.

Variable A : Type.

Inductive repeats : list A -> Prop :=            (*   repeats  *)
  repeats_hd l x : In x l    -> repeats (x::l)
| repeats_tl l x : repeats l -> repeats (x::l).

Fixpoint no_repeats (l: list A): Prop :=
  match l with nil => True | a::l' =>  ~ In a l' /\ no_repeats l' end.

Hypothesis not_no_repeats_repeats: forall l, (~ no_repeats l) -> repeats l.

Lemma eq_nn_elim (a b : A) : ~ a <> b -> a = b.
Proof.
  intros H.
  assert (H' : ~ no_repeats [a; b]).
  { simpl. intuition. }
  apply not_no_repeats_repeats in H'.
  inversion H'; subst.
  { subst. simpl in *. intuition; tauto. }
  inversion H1; simpl in *; subst; intuition.
  inversion H2.
Qed.

Not every type supports eq_nn_elim, which means that you can only prove not_no_repeats_repeats by placing additional hypotheses on A. It should suffice to assume that A has decidable equality; that is:

Hypothesis eq_dec a b : a = b \/ a <> b.


来源:https://stackoverflow.com/questions/47699961/equality-between-functional-and-inductive-definitions

标签
易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!